From 8159be8d9895dd8b9a128d2f1b371a2dd8233e4c Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 26 Dec 2024 06:54:11 +0000 Subject: [PATCH 001/189] feat(Algebra/Colimit/ModuleRing): generalize to Semiring/AddCommMonoid (#20212) and remove all unnecessary IsDirected assumptions. --- .../Algebra/Category/ModuleCat/Limits.lean | 6 +- Mathlib/Algebra/Colimit/ModuleRing.lean | 393 ++++++++---------- Mathlib/GroupTheory/Congruence/Defs.lean | 11 + .../TensorProduct/DirectLimit.lean | 20 +- 4 files changed, 206 insertions(+), 224 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Limits.lean index e2f482888b756..0f97bf8276775 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Limits.lean @@ -204,10 +204,10 @@ section DirectLimit open Module variable {ι : Type v} -variable [dec_ι : DecidableEq ι] [Preorder ι] +variable [DecidableEq ι] [Preorder ι] variable (G : ι → Type v) variable [∀ i, AddCommGroup (G i)] [∀ i, Module R (G i)] -variable (f : ∀ i j, i ≤ j → G i →ₗ[R] G j) [DirectedSystem G fun i j h => f i j h] +variable (f : ∀ i j, i ≤ j → G i →ₗ[R] G j) [DirectedSystem G fun i j h ↦ f i j h] /-- The diagram (in the sense of `CategoryTheory`) of an unbundled `directLimit` of modules. -/ @@ -241,7 +241,7 @@ def directLimitCocone : Cocone (directLimitDiagram G f) where /-- The unbundled `directLimit` of modules is a colimit in the sense of `CategoryTheory`. -/ @[simps] -def directLimitIsColimit [IsDirected ι (· ≤ ·)] : IsColimit (directLimitCocone G f) where +def directLimitIsColimit : IsColimit (directLimitCocone G f) where desc s := ofHom <| Module.DirectLimit.lift R ι G f (fun i => (s.ι.app i).hom) fun i j h x => by simp only [Functor.const_obj_obj] diff --git a/Mathlib/Algebra/Colimit/ModuleRing.lean b/Mathlib/Algebra/Colimit/ModuleRing.lean index 6268046e01ac7..5b66a75497403 100644 --- a/Mathlib/Algebra/Colimit/ModuleRing.lean +++ b/Mathlib/Algebra/Colimit/ModuleRing.lean @@ -5,7 +5,6 @@ Authors: Kenny Lau, Chris Hughes, Jujian Zhang -/ import Mathlib.Algebra.DirectSum.Module import Mathlib.Data.Finset.Order -import Mathlib.LinearAlgebra.Quotient.Basic import Mathlib.RingTheory.FreeCommRing import Mathlib.RingTheory.Ideal.Maps import Mathlib.RingTheory.Ideal.Quotient.Defs @@ -34,45 +33,50 @@ so as to make the operations (addition etc.) "computable". suppress_compilation -universe u v v' v'' w u₁ +variable {R : Type*} [Semiring R] {ι : Type*} [Preorder ι] {G : ι → Type*} open Submodule -variable {R : Type u} [Ring R] -variable {ι : Type v} -variable [Preorder ι] -variable (G : ι → Type w) - namespace Module -variable [∀ i, AddCommGroup (G i)] [∀ i, Module R (G i)] -variable {G} -variable (f : ∀ i j, i ≤ j → G i →ₗ[R] G j) - alias DirectedSystem.map_self := DirectedSystem.map_self' alias DirectedSystem.map_map := DirectedSystem.map_map' +variable [DecidableEq ι] +variable [∀ i, AddCommMonoid (G i)] [∀ i, Module R (G i)] (f : ∀ i j, i ≤ j → G i →ₗ[R] G j) + +/-- The relation on the direct sum that generates the additive congruence that defines the +colimit as a quotient. -/ +inductive DirectLimit.Eqv : DirectSum ι G → DirectSum ι G → Prop + | of_map {i j} (h : i ≤ j) (x : G i) : + Eqv (DirectSum.lof R ι G i x) (DirectSum.lof R ι G j <| f i j h x) + variable (G) /-- The direct limit of a directed system is the modules glued together along the maps. -/ -def DirectLimit [DecidableEq ι] : Type max v w := - DirectSum ι G ⧸ - (span R <| - { a | - ∃ (i j : _) (H : i ≤ j) (x : _), - DirectSum.lof R ι G i x - DirectSum.lof R ι G j (f i j H x) = a }) +def DirectLimit [DecidableEq ι] : Type _ := (addConGen <| DirectLimit.Eqv f).Quotient namespace DirectLimit section Basic -variable [DecidableEq ι] +instance addCommMonoid : AddCommMonoid (DirectLimit G f) := + AddCon.addCommMonoid _ -instance addCommGroup : AddCommGroup (DirectLimit G f) := - Quotient.addCommGroup _ +instance module : Module R (DirectLimit G f) where + smul r := AddCon.lift _ ((AddCon.mk' _).comp <| smulAddHom R _ r) <| + AddCon.addConGen_le fun x y ⟨_, _⟩ ↦ (AddCon.eq _).mpr <| by + simpa only [smulAddHom_apply, ← map_smul] using .of _ _ (.of_map _ _) + one_smul := by rintro ⟨⟩; exact congr_arg _ (one_smul _ _) + mul_smul _ _ := by rintro ⟨⟩; exact congr_arg _ (mul_smul _ _ _) + smul_zero _ := congr_arg _ (smul_zero _) + smul_add _ := by rintro ⟨⟩ ⟨⟩; exact congr_arg _ (smul_add _ _ _) + add_smul _ _ := by rintro ⟨⟩; exact congr_arg _ (add_smul _ _ _) + zero_smul := by rintro ⟨⟩; exact congr_arg _ (zero_smul _ _) -instance module : Module R (DirectLimit G f) := - Quotient.module _ +instance addCommGroup (G : ι → Type*) [∀ i, AddCommGroup (G i)] [∀ i, Module R (G i)] + (f : ∀ i j, i ≤ j → G i →ₗ[R] G j) : AddCommGroup (DirectLimit G f) := + inferInstanceAs (AddCommGroup <| AddCon.Quotient _) instance inhabited : Inhabited (DirectLimit G f) := ⟨0⟩ @@ -84,76 +88,79 @@ variable (R ι) /-- The canonical map from a component to the direct limit. -/ def of (i) : G i →ₗ[R] DirectLimit G f := - (mkQ _).comp <| DirectSum.lof R ι G i + .comp { __ := AddCon.mk' _, map_smul' := fun _ _ ↦ rfl } <| DirectSum.lof R ι G i variable {R ι G f} +theorem quotMk_of (i x) : Quot.mk _ (.of G i x) = of R ι G f i x := rfl + @[simp] theorem of_f {i j hij x} : of R ι G f j (f i j hij x) = of R ι G f i x := - Eq.symm <| (Submodule.Quotient.eq _).2 <| subset_span ⟨i, j, hij, x, rfl⟩ + (AddCon.eq _).mpr <| .symm <| .of _ _ (.of_map _ _) /-- Every element of the direct limit corresponds to some element in some component of the directed system. -/ theorem exists_of [Nonempty ι] [IsDirected ι (· ≤ ·)] (z : DirectLimit G f) : ∃ i x, of R ι G f i x = z := - Nonempty.elim (by infer_instance) fun ind : ι => - Quotient.inductionOn' z fun z => - DirectSum.induction_on z ⟨ind, 0, LinearMap.map_zero _⟩ (fun i x => ⟨i, x, rfl⟩) - fun p q ⟨i, x, ihx⟩ ⟨j, y, ihy⟩ => + Nonempty.elim (by infer_instance) fun ind : ι ↦ + Quotient.inductionOn' z fun z ↦ + DirectSum.induction_on z ⟨ind, 0, LinearMap.map_zero _⟩ (fun i x ↦ ⟨i, x, rfl⟩) + fun p q ⟨i, x, ihx⟩ ⟨j, y, ihy⟩ ↦ let ⟨k, hik, hjk⟩ := exists_ge_ge i j ⟨k, f i k hik x + f j k hjk y, by rw [LinearMap.map_add, of_f, of_f, ihx, ihy] rfl ⟩ +theorem exists_of₂ [Nonempty ι] [IsDirected ι (· ≤ ·)] (z w : DirectLimit G f) : + ∃ i x y, of R ι G f i x = z ∧ of R ι G f i y = w := + have ⟨i, x, hx⟩ := exists_of z + have ⟨j, y, hy⟩ := exists_of w + have ⟨k, hik, hjk⟩ := exists_ge_ge i j + ⟨k, f i k hik x, f j k hjk y, by rw [of_f, hx], by rw [of_f, hy]⟩ + @[elab_as_elim] protected theorem induction_on [Nonempty ι] [IsDirected ι (· ≤ ·)] {C : DirectLimit G f → Prop} (z : DirectLimit G f) (ih : ∀ i x, C (of R ι G f i x)) : C z := let ⟨i, x, h⟩ := exists_of z h ▸ ih i x -variable {P : Type u₁} [AddCommGroup P] [Module R P] +variable {P : Type*} [AddCommMonoid P] [Module R P] variable (R ι G f) in /-- The universal property of the direct limit: maps from the components to another module that respect the directed system structure (i.e. make some diagram commute) give rise to a unique map out of the direct limit. -/ def lift (g : ∀ i, G i →ₗ[R] P) (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) : - DirectLimit G f →ₗ[R] P := - liftQ _ (DirectSum.toModule R ι P g) - (span_le.2 fun a ⟨i, j, hij, x, hx⟩ => by - rw [← hx, SetLike.mem_coe, LinearMap.sub_mem_ker_iff, DirectSum.toModule_lof, - DirectSum.toModule_lof, Hg]) + DirectLimit G f →ₗ[R] P where + __ := AddCon.lift _ (DirectSum.toModule R ι P g) <| + AddCon.addConGen_le fun _ _ ⟨_, _⟩ ↦ by simpa using (Hg _ _ _ _).symm + map_smul' r := by rintro ⟨x⟩; exact map_smul (DirectSum.toModule R ι P g) r x variable (g : ∀ i, G i →ₗ[R] P) (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) @[simp] theorem lift_of {i} (x) : lift R ι G f g Hg (of R ι G f i x) = g i x := DirectSum.toModule_lof R _ _ -theorem lift_unique [IsDirected ι (· ≤ ·)] (F : DirectLimit G f →ₗ[R] P) (x) : - F x = - lift R ι G f (fun i => F.comp <| of R ι G f i) - (fun i j hij x => by rw [LinearMap.comp_apply, of_f]; rfl) x := by - cases isEmpty_or_nonempty ι - · simp_rw [Subsingleton.elim x 0, _root_.map_zero] - · exact DirectLimit.induction_on x fun i x => by rw [lift_of]; rfl +theorem lift_unique (F : DirectLimit G f →ₗ[R] P) (x) : + F x = lift R ι G f (fun i ↦ F.comp <| of R ι G f i) (fun i j hij x ↦ by simp) x := by + rcases x with ⟨x⟩ + exact x.induction_on (by simp) (fun _ _ ↦ .symm <| lift_of ..) (by simp +contextual) lemma lift_injective [IsDirected ι (· ≤ ·)] (injective : ∀ i, Function.Injective <| g i) : Function.Injective (lift R ι G f g Hg) := by cases isEmpty_or_nonempty ι · apply Function.injective_of_subsingleton - simp_rw [injective_iff_map_eq_zero] at injective ⊢ - intros z hz - induction z using DirectLimit.induction_on with - | ih _ g => - rw [lift_of] at hz - rw [injective _ g hz, _root_.map_zero] + intros z w eq + obtain ⟨i, x, y, rfl, rfl⟩ := exists_of₂ z w + simp_rw [lift_of] at eq + rw [injective _ eq] section functorial -variable {G' : ι → Type v'} [∀ i, AddCommGroup (G' i)] [∀ i, Module R (G' i)] +variable {G' : ι → Type*} [∀ i, AddCommMonoid (G' i)] [∀ i, Module R (G' i)] variable {f' : ∀ i j, i ≤ j → G' i →ₗ[R] G' j} -variable {G'' : ι → Type v''} [∀ i, AddCommGroup (G'' i)] [∀ i, Module R (G'' i)] +variable {G'' : ι → Type*} [∀ i, AddCommMonoid (G'' i)] [∀ i, Module R (G'' i)] variable {f'' : ∀ i j, i ≤ j → G'' i →ₗ[R] G'' j} /-- @@ -164,11 +171,9 @@ family of linear maps `gᵢ : Gᵢ ⟶ G'ᵢ` such that `g ∘ f = f' ∘ g` ind def map (g : (i : ι) → G i →ₗ[R] G' i) (hg : ∀ i j h, g j ∘ₗ f i j h = f' i j h ∘ₗ g i) : DirectLimit G f →ₗ[R] DirectLimit G' f' := lift _ _ _ _ (fun i ↦ of _ _ _ _ _ ∘ₗ g i) fun i j h g ↦ by - cases isEmpty_or_nonempty ι - · subsingleton - · have eq1 := LinearMap.congr_fun (hg i j h) g - simp only [LinearMap.coe_comp, Function.comp_apply] at eq1 ⊢ - rw [eq1, of_f] + have eq1 := LinearMap.congr_fun (hg i j h) g + simp only [LinearMap.coe_comp, Function.comp_apply] at eq1 ⊢ + rw [eq1, of_f] @[simp] lemma map_apply_of (g : (i : ι) → G i →ₗ[R] G' i) (hg : ∀ i j h, g j ∘ₗ f i j h = f' i j h ∘ₗ g i) @@ -176,13 +181,12 @@ def map (g : (i : ι) → G i →ₗ[R] G' i) (hg : ∀ i j h, g j ∘ₗ f i j map g hg (of _ _ G f _ x) = of R ι G' f' i (g i x) := lift_of _ _ _ -@[simp] lemma map_id [IsDirected ι (· ≤ ·)] : +@[simp] lemma map_id : map (fun _ ↦ LinearMap.id) (fun _ _ _ ↦ rfl) = LinearMap.id (R := R) (M := DirectLimit G f) := - DFunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (by subsingleton) fun _ ↦ - x.induction_on fun i g ↦ by simp + DFunLike.ext _ _ <| by + rintro ⟨x⟩; refine x.induction_on (by simp) (fun _ ↦ map_apply_of _ _) (by simp +contextual) -lemma map_comp [IsDirected ι (· ≤ ·)] - (g₁ : (i : ι) → G i →ₗ[R] G' i) (g₂ : (i : ι) → G' i →ₗ[R] G'' i) +lemma map_comp (g₁ : (i : ι) → G i →ₗ[R] G' i) (g₂ : (i : ι) → G' i →ₗ[R] G'' i) (hg₁ : ∀ i j h, g₁ j ∘ₗ f i j h = f' i j h ∘ₗ g₁ i) (hg₂ : ∀ i j h, g₂ j ∘ₗ f' i j h = f'' i j h ∘ₗ g₂ i) : (map g₂ hg₂ ∘ₗ map g₁ hg₁ : @@ -190,8 +194,10 @@ lemma map_comp [IsDirected ι (· ≤ ·)] (map (fun i ↦ g₂ i ∘ₗ g₁ i) fun i j h ↦ by rw [LinearMap.comp_assoc, hg₁ i, ← LinearMap.comp_assoc, hg₂ i, LinearMap.comp_assoc] : DirectLimit G f →ₗ[R] DirectLimit G'' f'') := - DFunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (by subsingleton) fun _ ↦ - x.induction_on fun i g ↦ by simp + DFunLike.ext _ _ <| by + rintro ⟨x⟩; refine x.induction_on (by simp) (fun _ _ ↦ ?_) (by simp +contextual) + show map g₂ hg₂ (map g₁ hg₁ <| of _ _ _ _ _ _) = map _ _ (of _ _ _ _ _ _) + simp_rw [map_apply_of]; rfl open LinearEquiv LinearMap in /-- @@ -199,8 +205,7 @@ Consider direct limits `lim G` and `lim G'` with direct system `f` and `f'` resp family of equivalences `eᵢ : Gᵢ ≅ G'ᵢ` such that `e ∘ f = f' ∘ e` induces an equivalence `lim G ≅ lim G'`. -/ -def congr [IsDirected ι (· ≤ ·)] - (e : (i : ι) → G i ≃ₗ[R] G' i) (he : ∀ i j h, e j ∘ₗ f i j h = f' i j h ∘ₗ e i) : +def congr (e : (i : ι) → G i ≃ₗ[R] G' i) (he : ∀ i j h, e j ∘ₗ f i j h = f' i j h ∘ₗ e i) : DirectLimit G f ≃ₗ[R] DirectLimit G' f' := LinearEquiv.ofLinear (map (e ·) he) (map (fun i ↦ (e i).symm) fun i j h ↦ by @@ -208,16 +213,14 @@ def congr [IsDirected ι (· ≤ ·)] refl_toLinearMap, comp_id]) (by simp [map_comp]) (by simp [map_comp]) -lemma congr_apply_of [IsDirected ι (· ≤ ·)] - (e : (i : ι) → G i ≃ₗ[R] G' i) (he : ∀ i j h, e j ∘ₗ f i j h = f' i j h ∘ₗ e i) +lemma congr_apply_of (e : (i : ι) → G i ≃ₗ[R] G' i) (he : ∀ i j h, e j ∘ₗ f i j h = f' i j h ∘ₗ e i) {i : ι} (g : G i) : congr e he (of _ _ G f i g) = of _ _ G' f' i (e i g) := map_apply_of _ he _ open LinearEquiv LinearMap in -lemma congr_symm_apply_of [IsDirected ι (· ≤ ·)] - (e : (i : ι) → G i ≃ₗ[R] G' i) (he : ∀ i j h, e j ∘ₗ f i j h = f' i j h ∘ₗ e i) - {i : ι} (g : G' i) : +lemma congr_symm_apply_of (e : (i : ι) → G i ≃ₗ[R] G' i) + (he : ∀ i j h, e j ∘ₗ f i j h = f' i j h ∘ₗ e i) {i : ι} (g : G' i) : (congr e he).symm (of _ _ G' f' i g) = of _ _ G f i ((e i).symm g) := map_apply_of _ (fun i j h ↦ by rw [toLinearMap_symm_comp_eq, ← comp_assoc, he i, comp_assoc, comp_coe, symm_trans_self, @@ -229,7 +232,7 @@ end Basic section equiv -variable [Nonempty ι] [IsDirected ι (· ≤ ·)] [DecidableEq ι] [DirectedSystem G (f · · ·)] +variable [Nonempty ι] [IsDirected ι (· ≤ ·)] [DirectedSystem G (f · · ·)] open _root_.DirectLimit /-- The direct limit constructed as a quotient of the direct sum is isomorphic to @@ -237,8 +240,9 @@ the direct limit constructed as a quotient of the disjoint union. -/ def linearEquiv : DirectLimit G f ≃ₗ[R] _root_.DirectLimit G f := .ofLinear (lift _ _ _ _ (Module.of _ _ _ _) fun _ _ _ _ ↦ .symm <| eq_of_le ..) (Module.lift _ _ _ _ (of _ _ _ _) fun _ _ _ _ ↦ of_f ..) - (by ext ⟨_⟩; rw [← Quotient.mk]; simp [Module.lift, _root_.DirectLimit.lift_def]; rfl) - (by ext x; refine x.induction_on fun i x ↦ ?_; simp; rfl) + (by ext ⟨_⟩; rw [← Quotient.mk]; simp [Module.lift, _root_.DirectLimit.lift_def]; rfl) <| by + ext ⟨x⟩; refine x.induction_on (by simp) (fun i x ↦ ?_) (by simp+contextual) + rw [quotMk_of, LinearMap.comp_apply, lift_of, Module.lift_of, LinearMap.id_apply] theorem linearEquiv_of {i g} : linearEquiv _ _ (of _ _ G f i g) = ⟦⟨i, g⟩⟧ := by simp [linearEquiv]; rfl @@ -251,7 +255,7 @@ variable {G f} /-- A component that corresponds to zero in the direct limit is already zero in some bigger module in the directed system. -/ -theorem of.zero_exact [DirectedSystem G (f · · ·)] [IsDirected ι (· ≤ ·)] [DecidableEq ι] +theorem of.zero_exact [DirectedSystem G (f · · ·)] [IsDirected ι (· ≤ ·)] {i x} (H : of R ι G f i x = 0) : ∃ j hij, f i j hij x = (0 : G j) := by haveI : Nonempty ι := ⟨i⟩ @@ -264,26 +268,28 @@ end Module namespace AddCommGroup -variable [∀ i, AddCommGroup (G i)] +variable (G) [∀ i, AddCommMonoid (G i)] /-- The direct limit of a directed system is the abelian groups glued together along the maps. -/ 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) _ + @Module.DirectLimit ℕ _ ι _ G _ _ (fun i j hij ↦ (f i j hij).toNatLinearMap) _ namespace DirectLimit variable (f : ∀ i j, i ≤ j → G i →+ G j) -protected theorem directedSystem [h : DirectedSystem G fun i j h => f i j h] : - DirectedSystem G fun i j hij => (f i j hij).toIntLinearMap := +local instance directedSystem [h : DirectedSystem G fun i j h ↦ f i j h] : + DirectedSystem G fun i j hij ↦ (f i j hij).toNatLinearMap := 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 +instance : AddCommMonoid (DirectLimit G f) := + Module.DirectLimit.addCommMonoid G fun i j hij ↦ (f i j hij).toNatLinearMap + +instance addCommGroup (G : ι → Type*) [∀ i, AddCommGroup (G i)] + (f : ∀ i j, i ≤ j → G i →+ G j) : AddCommGroup (DirectLimit G f) := + Module.DirectLimit.addCommGroup G fun i j hij ↦ (f i j hij).toNatLinearMap instance : Inhabited (DirectLimit G f) := ⟨0⟩ @@ -292,7 +298,7 @@ instance [IsEmpty ι] : Unique (DirectLimit G f) := Module.DirectLimit.unique _ /-- The canonical map from a component to the direct limit. -/ def of (i) : G i →+ DirectLimit G f := - (Module.DirectLimit.of ℤ ι G (fun i j hij => (f i j hij).toIntLinearMap) i).toAddMonoidHom + (Module.DirectLimit.of ℕ ι G _ i).toAddMonoidHom variable {G f} @@ -307,11 +313,11 @@ protected theorem induction_on [Nonempty ι] [IsDirected ι (· ≤ ·)] {C : Di /-- A component that corresponds to zero in the direct limit is already zero in some bigger module in the directed system. -/ -theorem of.zero_exact [IsDirected ι (· ≤ ·)] [DirectedSystem G fun i j h => f i j h] (i x) +theorem of.zero_exact [IsDirected ι (· ≤ ·)] [DirectedSystem G fun i j h ↦ f i j h] (i x) (h : of G f i x = 0) : ∃ j hij, f i j hij x = 0 := Module.DirectLimit.of.zero_exact h -variable (P : Type u₁) [AddCommGroup P] +variable (P : Type*) [AddCommMonoid P] variable (g : ∀ i, G i →+ P) variable (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) variable (G f) @@ -320,8 +326,8 @@ variable (G f) that respect the directed system structure (i.e. make some diagram commute) give rise to a unique map out of the direct limit. -/ def lift : DirectLimit G f →+ P := - (Module.DirectLimit.lift ℤ ι G (fun i j hij => (f i j hij).toIntLinearMap) - (fun i => (g i).toIntLinearMap) Hg).toAddMonoidHom + (Module.DirectLimit.lift ℕ ι G (fun i j hij ↦ (f i j hij).toNatLinearMap) + (fun i ↦ (g i).toNatLinearMap) Hg).toAddMonoidHom variable {G f} @@ -329,32 +335,26 @@ variable {G f} theorem lift_of (i x) : lift G f P g Hg (of G f i x) = g i x := Module.DirectLimit.lift_of -- Note: had to make these arguments explicit https://github.com/leanprover-community/mathlib4/pull/8386 - (f := (fun i j hij => (f i j hij).toIntLinearMap)) - (fun i => (g i).toIntLinearMap) + (f := fun i j hij ↦ (f i j hij).toNatLinearMap) + (fun i ↦ (g i).toNatLinearMap) Hg x -theorem lift_unique [IsDirected ι (· ≤ ·)] (F : DirectLimit G f →+ P) (x) : - F x = lift G f P (fun i => F.comp (of G f i)) (fun i j hij x => by simp) x := by - cases isEmpty_or_nonempty ι - · simp_rw [Subsingleton.elim x 0, _root_.map_zero] - · exact DirectLimit.induction_on x fun i x => by simp +theorem lift_unique (F : DirectLimit G f →+ P) (x) : + F x = lift G f P (fun i ↦ F.comp (of G f i)) (fun i j hij x ↦ by simp) x := by + rcases x with ⟨x⟩ + exact x.induction_on (by simp) (fun _ _ ↦ .symm <| lift_of ..) (by simp +contextual) lemma lift_injective [IsDirected ι (· ≤ ·)] (injective : ∀ i, Function.Injective <| g i) : - Function.Injective (lift G f P g Hg) := by - cases isEmpty_or_nonempty ι - · apply Function.injective_of_subsingleton - simp_rw [injective_iff_map_eq_zero] at injective ⊢ - intros z hz - induction z using DirectLimit.induction_on with - | ih _ g => rw [lift_of] at hz; rw [injective _ g hz, _root_.map_zero] + Function.Injective (lift G f P g Hg) := + Module.DirectLimit.lift_injective (f := fun i j hij ↦ (f i j hij).toNatLinearMap) _ Hg injective section functorial -variable {G' : ι → Type v'} [∀ i, AddCommGroup (G' i)] +variable {G' : ι → Type*} [∀ i, AddCommMonoid (G' i)] variable {f' : ∀ i j, i ≤ j → G' i →+ G' j} -variable {G'' : ι → Type v''} [∀ i, AddCommGroup (G'' i)] +variable {G'' : ι → Type*} [∀ i, AddCommMonoid (G'' i)] variable {f'' : ∀ i j, i ≤ j → G'' i →+ G'' j} /-- @@ -366,11 +366,9 @@ def map (g : (i : ι) → G i →+ G' i) (hg : ∀ i j h, (g j).comp (f i j h) = (f' i j h).comp (g i)) : DirectLimit G f →+ DirectLimit G' f' := lift _ _ _ (fun i ↦ (of _ _ _).comp (g i)) fun i j h g ↦ by - cases isEmpty_or_nonempty ι - · subsingleton - · have eq1 := DFunLike.congr_fun (hg i j h) g - simp only [AddMonoidHom.coe_comp, Function.comp_apply] at eq1 ⊢ - rw [eq1, of_f] + have eq1 := DFunLike.congr_fun (hg i j h) g + simp only [AddMonoidHom.coe_comp, Function.comp_apply] at eq1 ⊢ + rw [eq1, of_f] @[simp] lemma map_apply_of (g : (i : ι) → G i →+ G' i) (hg : ∀ i j h, (g j).comp (f i j h) = (f' i j h).comp (g i)) @@ -378,13 +376,12 @@ def map (g : (i : ι) → G i →+ G' i) map g hg (of G f _ x) = of G' f' i (g i x) := lift_of _ _ _ _ _ -@[simp] lemma map_id [IsDirected ι (· ≤ ·)] : +@[simp] lemma map_id : map (fun _ ↦ AddMonoidHom.id _) (fun _ _ _ ↦ rfl) = AddMonoidHom.id (DirectLimit G f) := - DFunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (by subsingleton) fun _ ↦ - x.induction_on fun i g ↦ by simp + DFunLike.ext _ _ <| by + rintro ⟨x⟩; refine x.induction_on (by simp) (fun _ ↦ map_apply_of _ _) (by simp +contextual) -lemma map_comp [IsDirected ι (· ≤ ·)] - (g₁ : (i : ι) → G i →+ G' i) (g₂ : (i : ι) → G' i →+ G'' i) +lemma map_comp (g₁ : (i : ι) → G i →+ G' i) (g₂ : (i : ι) → G' i →+ G'' i) (hg₁ : ∀ i j h, (g₁ j).comp (f i j h) = (f' i j h).comp (g₁ i)) (hg₂ : ∀ i j h, (g₂ j).comp (f' i j h) = (f'' i j h).comp (g₂ i)) : ((map g₂ hg₂).comp (map g₁ hg₁) : @@ -393,16 +390,17 @@ lemma map_comp [IsDirected ι (· ≤ ·)] rw [AddMonoidHom.comp_assoc, hg₁ i, ← AddMonoidHom.comp_assoc, hg₂ i, AddMonoidHom.comp_assoc] : DirectLimit G f →+ DirectLimit G'' f'') := - DFunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (by subsingleton) fun _ ↦ - x.induction_on fun i g ↦ by simp + DFunLike.ext _ _ <| by + rintro ⟨x⟩; refine x.induction_on (by simp) (fun _ _ ↦ ?_) (by simp +contextual) + show map g₂ hg₂ (map g₁ hg₁ <| of _ _ _ _) = map _ _ (of _ _ _ _) + simp_rw [map_apply_of]; rfl /-- Consider direct limits `lim G` and `lim G'` with direct system `f` and `f'` respectively, any family of equivalences `eᵢ : Gᵢ ≅ G'ᵢ` such that `e ∘ f = f' ∘ e` induces an equivalence `lim G ⟶ lim G'`. -/ -def congr [IsDirected ι (· ≤ ·)] - (e : (i : ι) → G i ≃+ G' i) +def congr (e : (i : ι) → G i ≃+ G' i) (he : ∀ i j h, (e j).toAddMonoidHom.comp (f i j h) = (f' i j h).comp (e i)) : DirectLimit G f ≃+ DirectLimit G' f' := AddMonoidHom.toAddEquiv (map (e ·) he) @@ -411,18 +409,15 @@ def congr [IsDirected ι (· ≤ ·)] simp only [AddMonoidHom.coe_comp, AddEquiv.coe_toAddMonoidHom, Function.comp_apply, AddMonoidHom.coe_coe, AddEquiv.apply_symm_apply] at eq1 ⊢ simp [← eq1, of_f]) - (by rw [map_comp]; convert map_id <;> aesop) - (by rw [map_comp]; convert map_id <;> aesop) + (by simp [map_comp]) (by simp [map_comp]) -lemma congr_apply_of [IsDirected ι (· ≤ ·)] - (e : (i : ι) → G i ≃+ G' i) +lemma congr_apply_of (e : (i : ι) → G i ≃+ G' i) (he : ∀ i j h, (e j).toAddMonoidHom.comp (f i j h) = (f' i j h).comp (e i)) {i : ι} (g : G i) : congr e he (of G f i g) = of G' f' i (e i g) := map_apply_of _ he _ -lemma congr_symm_apply_of [IsDirected ι (· ≤ ·)] - (e : (i : ι) → G i ≃+ G' i) +lemma congr_symm_apply_of (e : (i : ι) → G i ≃+ G' i) (he : ∀ i j h, (e j).toAddMonoidHom.comp (f i j h) = (f' i j h).comp (e i)) {i : ι} (g : G' i) : (congr e he).symm (of G' f' i g) = of G f i ((e i).symm g) := by @@ -436,7 +431,7 @@ end AddCommGroup namespace Ring -variable [∀ i, CommRing (G i)] +variable (G) [∀ i, CommRing (G i)] section @@ -445,14 +440,14 @@ variable (f : ∀ i j, i ≤ j → G i → G j) open FreeCommRing /-- The direct limit of a directed system is the rings glued together along the maps. -/ -def DirectLimit : Type max v w := - FreeCommRing (Σi, G i) ⧸ +def DirectLimit : Type _ := + FreeCommRing (Σ i, G i) ⧸ Ideal.span { a | - (∃ i j H x, of (⟨j, f i j H x⟩ : Σi, G i) - of ⟨i, x⟩ = a) ∨ - (∃ i, of (⟨i, 1⟩ : Σi, G i) - 1 = a) ∨ - (∃ i x y, of (⟨i, x + y⟩ : Σi, G i) - (of ⟨i, x⟩ + of ⟨i, y⟩) = a) ∨ - ∃ i x y, of (⟨i, x * y⟩ : Σi, G i) - of ⟨i, x⟩ * of ⟨i, y⟩ = a } + (∃ i j H x, of (⟨j, f i j H x⟩ : Σ i, G i) - of ⟨i, x⟩ = a) ∨ + (∃ i, of (⟨i, 1⟩ : Σ i, G i) - 1 = a) ∨ + (∃ i x y, of (⟨i, x + y⟩ : Σ i, G i) - (of ⟨i, x⟩ + of ⟨i, y⟩) = a) ∨ + ∃ i x y, of (⟨i, x * y⟩ : Σ i, G i) - of ⟨i, x⟩ * of ⟨i, y⟩ = a } namespace DirectLimit @@ -473,42 +468,29 @@ instance : Inhabited (DirectLimit G f) := /-- The canonical map from a component to the direct limit. -/ nonrec def of (i) : G i →+* DirectLimit G f := RingHom.mk' - { toFun := fun x => Ideal.Quotient.mk _ (of (⟨i, x⟩ : Σi, G i)) + { toFun := fun x ↦ Ideal.Quotient.mk _ (of (⟨i, x⟩ : Σ i, G i)) map_one' := Ideal.Quotient.eq.2 <| subset_span <| Or.inr <| Or.inl ⟨i, rfl⟩ - map_mul' := fun x y => + map_mul' := fun x y ↦ Ideal.Quotient.eq.2 <| subset_span <| Or.inr <| Or.inr <| Or.inr ⟨i, x, y, rfl⟩ } - fun x y => Ideal.Quotient.eq.2 <| subset_span <| Or.inr <| Or.inr <| Or.inl ⟨i, x, y, rfl⟩ + fun x y ↦ Ideal.Quotient.eq.2 <| subset_span <| Or.inr <| Or.inr <| Or.inl ⟨i, x, y, rfl⟩ variable {G f} +theorem quotientMk_of (i x) : Ideal.Quotient.mk _ (.of ⟨i, x⟩) = of G f i x := + rfl + @[simp] theorem of_f {i j} (hij) (x) : of G f j (f i j hij x) = of G f i x := Ideal.Quotient.eq.2 <| subset_span <| Or.inl ⟨i, j, hij, x, rfl⟩ /-- Every element of the direct limit corresponds to some element in some component of the directed system. -/ theorem exists_of [Nonempty ι] [IsDirected ι (· ≤ ·)] (z : DirectLimit G f) : - ∃ i x, of G f i x = z := - Nonempty.elim (by infer_instance) fun ind : ι => - Quotient.inductionOn' z fun x => - FreeAbelianGroup.induction_on x ⟨ind, 0, (of _ _ ind).map_zero⟩ - (fun s => - Multiset.induction_on s ⟨ind, 1, (of _ _ ind).map_one⟩ fun a s ih => - let ⟨i, x⟩ := a - let ⟨j, y, hs⟩ := ih - let ⟨k, hik, hjk⟩ := exists_ge_ge i j - ⟨k, f i k hik x * f j k hjk y, by - rw [(of G f k).map_mul, of_f, of_f, hs] - /- porting note: In Lean3, from here, this was `by refl`. I have added - the lemma `FreeCommRing.of_cons` to fix this proof. -/ - apply congr_arg Quotient.mk'' - symm - apply FreeCommRing.of_cons⟩) - (fun s ⟨i, x, ih⟩ => ⟨i, -x, by - rw [(of G _ _).map_neg, ih] - rfl⟩) - fun p q ⟨i, x, ihx⟩ ⟨j, y, ihy⟩ => - let ⟨k, hik, hjk⟩ := exists_ge_ge i j - ⟨k, f i k hik x + f j k hjk y, by rw [(of _ _ _).map_add, of_f, of_f, ihx, ihy]; rfl⟩ + ∃ i x, of G f i x = z := by + obtain ⟨z, rfl⟩ := Ideal.Quotient.mk_surjective z + refine z.induction_on ⟨Classical.arbitrary ι, -1, by simp⟩ (fun ⟨i, x⟩ ↦ ⟨i, x, rfl⟩) ?_ ?_ <;> + rintro x' y' ⟨i, x, hx⟩ ⟨j, y, hy⟩ <;> have ⟨k, hik, hjk⟩ := exists_ge_ge i j + · exact ⟨k, f i k hik x + f j k hjk y, by rw [map_add, of_f, of_f, hx, hy]; rfl⟩ + · exact ⟨k, f i k hik x * f j k hjk y, by rw [map_mul, of_f, of_f, hx, hy]; rfl⟩ section @@ -517,18 +499,18 @@ open Polynomial variable {f' : ∀ i j, i ≤ j → G i →+* G j} nonrec theorem Polynomial.exists_of [Nonempty ι] [IsDirected ι (· ≤ ·)] - (q : Polynomial (DirectLimit G fun i j h => f' i j h)) : - ∃ i p, Polynomial.map (of G (fun i j h => f' i j h) i) p = q := + (q : Polynomial (DirectLimit G fun i j h ↦ f' i j h)) : + ∃ i p, Polynomial.map (of G (fun i j h ↦ f' i j h) i) p = q := Polynomial.induction_on q - (fun z => + (fun z ↦ let ⟨i, x, h⟩ := exists_of z ⟨i, C x, by rw [map_C, h]⟩) - (fun q₁ q₂ ⟨i₁, p₁, ih₁⟩ ⟨i₂, p₂, ih₂⟩ => + (fun q₁ q₂ ⟨i₁, p₁, ih₁⟩ ⟨i₂, p₂, ih₂⟩ ↦ let ⟨i, h1, h2⟩ := exists_ge_ge i₁ i₂ ⟨i, p₁.map (f' i₁ i h1) + p₂.map (f' i₂ i h2), by rw [Polynomial.map_add, map_map, map_map, ← ih₁, ← ih₂] congr 2 <;> ext x <;> simp_rw [RingHom.comp_apply, of_f]⟩) - fun n z _ => + fun n z _ ↦ let ⟨i, x, h⟩ := exists_of z ⟨i, C x * X ^ (n + 1), by rw [Polynomial.map_mul, map_C, h, Polynomial.map_pow, map_X]⟩ @@ -540,7 +522,7 @@ theorem induction_on [Nonempty ι] [IsDirected ι (· ≤ ·)] {C : DirectLimit let ⟨i, x, hx⟩ := exists_of z hx ▸ ih i x -variable (P : Type u₁) [CommRing P] +variable (P : Type*) [CommRing P] open FreeCommRing @@ -551,10 +533,10 @@ to a unique map out of the direct limit. -/ def lift (g : ∀ i, G i →+* P) (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) : DirectLimit G f →+* P := - Ideal.Quotient.lift _ (FreeCommRing.lift fun x : Σi, G i => g x.1 x.2) + Ideal.Quotient.lift _ (FreeCommRing.lift fun x : Σ i, G i ↦ g x.1 x.2) (by suffices Ideal.span _ ≤ - Ideal.comap (FreeCommRing.lift fun x : Σi : ι, G i => g x.fst x.snd) ⊥ by + Ideal.comap (FreeCommRing.lift fun x : Σ i : ι, G i ↦ g x.fst x.snd) ⊥ by intro x hx exact (mem_bot P).1 (this hx) rw [Ideal.span_le] @@ -569,14 +551,11 @@ variable (g : ∀ i, G i →+* P) (Hg : ∀ i j hij x, g j (f i j hij x) = g i x @[simp] theorem lift_of (i x) : lift G f P g Hg (of G f i x) = g i x := FreeCommRing.lift_of _ _ -theorem lift_unique [IsDirected ι (· ≤ ·)] (F : DirectLimit G f →+* P) (x) : - F x = lift G f P (fun i => F.comp <| of G f i) (fun i j hij x => by simp [of_f]) x := by - cases isEmpty_or_nonempty ι - · apply DFunLike.congr_fun - apply Ideal.Quotient.ringHom_ext - refine FreeCommRing.hom_ext fun ⟨i, _⟩ ↦ ?_ - exact IsEmpty.elim' inferInstance i - · exact DirectLimit.induction_on x fun i x => by simp +theorem lift_unique (F : DirectLimit G f →+* P) (x) : + F x = lift G f P (fun i ↦ F.comp <| of G f i) (fun i j hij x ↦ by simp) x := by + obtain ⟨x, rfl⟩ := Ideal.Quotient.mk_surjective x + exact x.induction_on (by simp) (fun _ ↦ .symm <| lift_of ..) + (by simp +contextual) (by simp+contextual) lemma lift_injective [Nonempty ι] [IsDirected ι (· ≤ ·)] (injective : ∀ i, Function.Injective <| g i) : @@ -589,7 +568,7 @@ lemma lift_injective [Nonempty ι] [IsDirected ι (· ≤ ·)] section OfZeroExact variable (f' : ∀ i j, i ≤ j → G i →+* G j) -variable [DirectedSystem G fun i j h => f' i j h] [IsDirected ι (· ≤ ·)] +variable [DirectedSystem G fun i j h ↦ f' i j h] [IsDirected ι (· ≤ ·)] variable (G f) open _root_.DirectLimit in @@ -611,7 +590,7 @@ variable {G f'} bigger module in the directed system. -/ theorem of.zero_exact {i x} (hix : of G (f' · · ·) i x = 0) : ∃ (j : _) (hij : i ≤ j), f' i j hij x = 0 := by - haveI : Nonempty ι := ⟨i⟩ + haveI := Nonempty.intro i apply_fun ringEquiv _ _ at hix rwa [map_zero, ringEquiv_of, DirectLimit.exists_eq_zero] at hix @@ -621,25 +600,19 @@ variable (f' : ∀ i j, i ≤ j → G i →+* G j) /-- If the maps in the directed system are injective, then the canonical maps from the components to the direct limits are injective. -/ -theorem of_injective [IsDirected ι (· ≤ ·)] [DirectedSystem G fun i j h => f' i j h] +theorem of_injective [IsDirected ι (· ≤ ·)] [DirectedSystem G fun i j h ↦ f' i j h] (hf : ∀ i j hij, Function.Injective (f' i j hij)) (i) : - Function.Injective (of G (fun i j h => f' i j h) i) := by - suffices ∀ x, of G (fun i j h => f' i j h) i x = 0 → x = 0 by - intro x y hxy - rw [← sub_eq_zero] - apply this - rw [(of G _ i).map_sub, hxy, sub_self] - intro x hx - rcases of.zero_exact hx with ⟨j, hij, hfx⟩ - apply hf i j hij - rw [hfx, (f' i j hij).map_zero] + Function.Injective (of G (fun i j h ↦ f' i j h) i) := + haveI := Nonempty.intro i + ((ringEquiv _ _).comp_injective _).mp + fun _ _ eq ↦ DirectLimit.mk_injective f' hf _ (by simpa only [← ringEquiv_of]) section functorial variable {f : ∀ i j, i ≤ j → G i →+* G j} -variable {G' : ι → Type v'} [∀ i, CommRing (G' i)] +variable {G' : ι → Type*} [∀ i, CommRing (G' i)] variable {f' : ∀ i j, i ≤ j → G' i →+* G' j} -variable {G'' : ι → Type v''} [∀ i, CommRing (G'' i)] +variable {G'' : ι → Type*} [∀ i, CommRing (G'' i)] variable {f'' : ∀ i j, i ≤ j → G'' i →+* G'' j} /-- @@ -661,15 +634,14 @@ 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 : + map (fun _ ↦ RingHom.id _) (fun _ _ _ ↦ rfl) = RingHom.id (DirectLimit G fun _ _ h ↦ f _ _ h) := + DFunLike.ext _ _ fun x ↦ by + obtain ⟨x, rfl⟩ := Ideal.Quotient.mk_surjective x + refine x.induction_on (by simp) (fun _ ↦ ?_) (by simp+contextual) (by simp+contextual) + rw [quotientMk_of, map_apply_of]; rfl -@[simp] lemma map_id [IsDirected ι (· ≤ ·)] : - map (fun _ ↦ RingHom.id _) (fun _ _ _ ↦ rfl) = - RingHom.id (DirectLimit G fun _ _ h ↦ f _ _ h) := - DFunLike.ext _ _ fun x ↦ x.induction_on fun i g ↦ by simp - -lemma map_comp [IsDirected ι (· ≤ ·)] - (g₁ : (i : ι) → G i →+* G' i) (g₂ : (i : ι) → G' i →+* G'' i) +lemma map_comp (g₁ : (i : ι) → G i →+* G' i) (g₂ : (i : ι) → G' i →+* G'' i) (hg₁ : ∀ i j h, (g₁ j).comp (f i j h) = (f' i j h).comp (g₁ i)) (hg₂ : ∀ i j h, (g₂ j).comp (f' i j h) = (f'' i j h).comp (g₂ i)) : ((map g₂ hg₂).comp (map g₁ hg₁) : @@ -677,40 +649,41 @@ lemma map_comp [IsDirected ι (· ≤ ·)] (map (fun i ↦ (g₂ i).comp (g₁ i)) fun i j h ↦ by rw [RingHom.comp_assoc, hg₁ i, ← RingHom.comp_assoc, hg₂ i, RingHom.comp_assoc] : DirectLimit G (fun _ _ h ↦ f _ _ h) →+* DirectLimit G'' fun _ _ h ↦ f'' _ _ h) := - DFunLike.ext _ _ fun x ↦ x.induction_on fun i g ↦ by simp + DFunLike.ext _ _ fun x ↦ by + obtain ⟨x, rfl⟩ := Ideal.Quotient.mk_surjective x + refine x.induction_on (by simp) (fun _ ↦ ?_) (by simp+contextual) (by simp+contextual) + rw [RingHom.comp_apply, quotientMk_of] + simp_rw [map_apply_of] + rfl /-- Consider direct limits `lim G` and `lim G'` with direct system `f` and `f'` respectively, any family of equivalences `eᵢ : Gᵢ ≅ G'ᵢ` such that `e ∘ f = f' ∘ e` induces an equivalence `lim G ⟶ lim G'`. -/ -def congr [IsDirected ι (· ≤ ·)] - (e : (i : ι) → G i ≃+* G' i) +def congr (e : (i : ι) → G i ≃+* G' i) (he : ∀ i j h, (e j).toRingHom.comp (f i j h) = (f' i j h).comp (e i)) : DirectLimit G (fun _ _ h ↦ f _ _ h) ≃+* DirectLimit G' fun _ _ h ↦ f' _ _ h := - RingEquiv.ofHomInv + RingEquiv.ofRingHom (map (e ·) he) (map (fun i ↦ (e i).symm) fun i j h ↦ DFunLike.ext _ _ fun x ↦ by have eq1 := DFunLike.congr_fun (he i j h) ((e i).symm x) simp only [RingEquiv.toRingHom_eq_coe, RingHom.coe_comp, RingHom.coe_coe, Function.comp_apply, RingEquiv.apply_symm_apply] at eq1 ⊢ simp [← eq1, of_f]) - (DFunLike.ext _ _ fun x ↦ x.induction_on <| by simp) - (DFunLike.ext _ _ fun x ↦ x.induction_on <| by simp) + (by simp [map_comp]) (by simp [map_comp]) -lemma congr_apply_of [IsDirected ι (· ≤ ·)] - (e : (i : ι) → G i ≃+* G' i) +lemma congr_apply_of (e : (i : ι) → G i ≃+* G' i) (he : ∀ i j h, (e j).toRingHom.comp (f i j h) = (f' i j h).comp (e i)) {i : ι} (g : G i) : congr e he (of G _ i g) = of G' (fun _ _ h ↦ f' _ _ h) i (e i g) := map_apply_of _ he _ -lemma congr_symm_apply_of [IsDirected ι (· ≤ ·)] - (e : (i : ι) → G i ≃+* G' i) +lemma congr_symm_apply_of (e : (i : ι) → G i ≃+* G' i) (he : ∀ i j h, (e j).toRingHom.comp (f i j h) = (f' i j h).comp (e i)) {i : ι} (g : G' i) : (congr e he).symm (of G' _ i g) = of G (fun _ _ h ↦ f _ _ h) i ((e i).symm g) := by - simp only [congr, RingEquiv.ofHomInv_symm_apply, map_apply_of, RingHom.coe_coe] + simp only [congr, RingEquiv.ofRingHom_symm_apply, map_apply_of, RingHom.coe_coe] end functorial @@ -722,7 +695,7 @@ end Ring namespace Field -variable [Nonempty ι] [IsDirected ι (· ≤ ·)] [∀ i, Field (G i)] +variable [Nonempty ι] [IsDirected ι (· ≤ ·)] [∀ i, Field (G i)] (G) variable (f : ∀ i j, i ≤ j → G i → G j) variable (f' : ∀ i j, i ≤ j → G i →+* G j) @@ -731,7 +704,7 @@ namespace DirectLimit instance nontrivial [DirectedSystem G (f' · · ·)] : Nontrivial (Ring.DirectLimit G (f' · · ·)) := ⟨⟨0, 1, - Nonempty.elim (by infer_instance) fun i : ι => by + Nonempty.elim (by infer_instance) fun i : ι ↦ by change (0 : Ring.DirectLimit G (f' · · ·)) ≠ 1 rw [← (Ring.DirectLimit.of _ _ _).map_one] · intro H; rcases Ring.DirectLimit.of.zero_exact H.symm with ⟨j, hij, hf⟩ @@ -739,10 +712,10 @@ instance nontrivial [DirectedSystem G (f' · · ·)] : exact one_ne_zero hf⟩⟩ 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.induction_on p fun i x H ↦ ⟨Ring.DirectLimit.of G f i x⁻¹, by rw [← (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 @@ -766,12 +739,12 @@ protected noncomputable abbrev field [DirectedSystem G (f' · · ·)] : -- This used to include the parent CommRing and Nontrivial instances, -- but leaving them implicit avoids a very expensive (2-3 minutes!) eta expansion. inv := inv G (f' · · ·) - mul_inv_cancel := fun _ => DirectLimit.mul_inv_cancel G (f' · · ·) + mul_inv_cancel := fun _ ↦ DirectLimit.mul_inv_cancel G (f' · · ·) inv_zero := dif_pos rfl nnqsmul := _ - nnqsmul_def := fun _ _ => rfl + nnqsmul_def _ _ := rfl qsmul := _ - qsmul_def := fun _ _ => rfl + qsmul_def _ _ := rfl end diff --git a/Mathlib/GroupTheory/Congruence/Defs.lean b/Mathlib/GroupTheory/Congruence/Defs.lean index 9ecaf70c760a9..909a38db18909 100644 --- a/Mathlib/GroupTheory/Congruence/Defs.lean +++ b/Mathlib/GroupTheory/Congruence/Defs.lean @@ -764,6 +764,17 @@ instance group : Group c.Quotient := toInv := Con.hasInv _ toDiv := Con.hasDiv _ } +/-- The quotient of a `CommGroup` by a congruence relation is a `CommGroup`. -/ +@[to_additive "The quotient of an `AddCommGroup` by an additive congruence +relation is an `AddCommGroup`."] +instance commGroup {M : Type*} [CommGroup M] (c : Con M) : CommGroup c.Quotient := + { (Function.Surjective.commGroup _ Quotient.mk''_surjective rfl (fun _ _ => rfl) (fun _ => rfl) + (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) : CommGroup c.Quotient) with + /- The `toGroup` field is given explicitly for performance reasons. + This avoids any need to unfold `Function.Surjective.commGroup` when the type checker is + checking that instance diagrams commute -/ + toGroup := Con.group _ } + end Groups section Units diff --git a/Mathlib/LinearAlgebra/TensorProduct/DirectLimit.lean b/Mathlib/LinearAlgebra/TensorProduct/DirectLimit.lean index d66c649d5ed79..bc41c773a8af6 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/DirectLimit.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/DirectLimit.lean @@ -22,13 +22,13 @@ as `R`-modules. open TensorProduct Module Module.DirectLimit -variable {R : Type*} [CommRing R] +variable {R : Type*} [CommSemiring R] variable {ι : Type*} variable [DecidableEq ι] [Preorder ι] variable {G : ι → Type*} -variable [∀ i, AddCommGroup (G i)] [∀ i, Module R (G i)] +variable [∀ i, AddCommMonoid (G i)] [∀ i, Module R (G i)] variable (f : ∀ i j, i ≤ j → G i →ₗ[R] G j) -variable (M : Type*) [AddCommGroup M] [Module R M] +variable (M : Type*) [AddCommMonoid M] [Module R M] -- alluding to the notation in `CategoryTheory.Monoidal` local notation M " ◁ " f => fun i j h ↦ LinearMap.lTensor M (f _ _ h) @@ -68,19 +68,17 @@ variable {M} in rw [toDirectLimit, lift.tmul, lift_of] rfl -variable [IsDirected ι (· ≤ ·)] - /-- `limᵢ (Gᵢ ⊗ M)` and `(limᵢ Gᵢ) ⊗ M` are isomorphic as modules -/ noncomputable def directLimitLeft : DirectLimit G f ⊗[R] M ≃ₗ[R] DirectLimit (G · ⊗[R] M) (f ▷ M) := by - refine LinearEquiv.ofLinear (toDirectLimit f M) (fromDirectLimit f M) ?_ ?_ - <;> cases isEmpty_or_nonempty ι - · ext; subsingleton - · refine DFunLike.ext _ _ fun x ↦ x.induction_on fun i g ↦ g.induction_on ?_ ?_ ?_ <;> aesop - · ext; subsingleton - · exact ext (DFunLike.ext _ _ fun g ↦ DFunLike.ext _ _ fun _ ↦ g.induction_on <| by aesop) + refine LinearEquiv.ofLinear (toDirectLimit f M) (fromDirectLimit f M) ?_ (ext ?_) + · ext ⟨x⟩ + exact x.induction_on (by simp) (fun i x ↦ x.induction_on (by simp) + (fun _ _ ↦ by rw [quotMk_of]; simp) <| by simp+contextual) (by simp+contextual) + · ext ⟨x⟩ m + exact x.induction_on (by simp) (fun _ _ ↦ by rw [quotMk_of]; simp) (by simp+contextual) @[simp] lemma directLimitLeft_tmul_of {i : ι} (g : G i) (m : M) : directLimitLeft f M (of _ _ _ _ _ g ⊗ₜ m) = of _ _ _ (f ▷ M) _ (g ⊗ₜ m) := From 4ccace04052e6463b8f46bf9aa98e8335aa7a5be Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 26 Dec 2024 07:24:13 +0000 Subject: [PATCH 002/189] chore(*): golf using `omega` (#20231) Found by https://github.com/dwrensha/tryAtEachStep --- Mathlib/Algebra/CubicDiscriminant.lean | 2 +- Mathlib/Algebra/GeomSum.lean | 11 +++---- Mathlib/Algebra/Homology/ExactSequence.lean | 2 +- .../Homology/HomotopyCategory/HomComplex.lean | 3 +- Mathlib/Algebra/Polynomial/Module/Basic.lean | 4 +-- Mathlib/Algebra/Polynomial/UnitTrinomial.lean | 10 ++---- Mathlib/AlgebraicTopology/DoldKan/Faces.lean | 5 +-- Mathlib/Analysis/Convex/Between.lean | 5 +-- .../SpecialFunctions/Trigonometric/Basic.lean | 2 +- Mathlib/CategoryTheory/ComposableArrows.lean | 16 +++------- Mathlib/CategoryTheory/Preadditive/Schur.lean | 5 +-- .../Combinatorics/Enumerative/Catalan.lean | 2 +- .../Combinatorics/Enumerative/DyckWord.lean | 3 +- Mathlib/Combinatorics/SimpleGraph/Walk.lean | 5 ++- Mathlib/Data/FP/Basic.lean | 10 +++--- Mathlib/Data/Int/Defs.lean | 16 ++-------- Mathlib/Data/List/NodupEquivFin.lean | 3 +- Mathlib/Data/List/Rotate.lean | 4 +-- Mathlib/Data/Nat/Defs.lean | 31 +++++-------------- Mathlib/Data/Nat/Squarefree.lean | 7 +---- Mathlib/Data/Ordmap/Ordset.lean | 5 +-- Mathlib/Data/Real/Archimedean.lean | 3 +- Mathlib/Data/ZMod/Basic.lean | 16 ++-------- Mathlib/Geometry/Euclidean/MongePoint.lean | 8 ++--- Mathlib/GroupTheory/Coxeter/Length.lean | 17 +++------- Mathlib/GroupTheory/FreeGroup/Basic.lean | 12 +------ Mathlib/GroupTheory/FreeGroup/Reduce.lean | 1 - .../SpecificGroups/Alternating.lean | 2 +- .../BilinearForm/Orthogonal.lean | 2 -- Mathlib/LinearAlgebra/Reflection.lean | 10 ++---- Mathlib/ModelTheory/Semantics.lean | 12 ++----- .../LegendreSymbol/QuadraticChar/Basic.lean | 4 +-- .../LegendreSymbol/QuadraticReciprocity.lean | 12 ++----- .../NumberTheory/Padics/PadicVal/Basic.lean | 8 +---- Mathlib/NumberTheory/Pell.lean | 2 +- Mathlib/NumberTheory/PellMatiyasevic.lean | 8 ++--- Mathlib/NumberTheory/PythagoreanTriples.lean | 8 ++--- Mathlib/NumberTheory/SumTwoSquares.lean | 2 +- .../Transcendental/Liouville/Residual.lean | 2 +- .../Zsqrtd/QuadraticReciprocity.lean | 7 +---- Mathlib/Order/RelSeries.lean | 8 ++--- .../RingTheory/MvPolynomial/Homogeneous.lean | 2 +- Mathlib/RingTheory/Polynomial/Bernstein.lean | 4 +-- Mathlib/RingTheory/PowerSeries/Basic.lean | 2 +- 44 files changed, 77 insertions(+), 226 deletions(-) diff --git a/Mathlib/Algebra/CubicDiscriminant.lean b/Mathlib/Algebra/CubicDiscriminant.lean index 04cd0d3ed3474..62b2934918c6c 100644 --- a/Mathlib/Algebra/CubicDiscriminant.lean +++ b/Mathlib/Algebra/CubicDiscriminant.lean @@ -80,7 +80,7 @@ private theorem coeffs : (∀ n > 3, P.toPoly.coeff n = 0) ∧ P.toPoly.coeff 3 norm_num intro n hn repeat' rw [if_neg] - any_goals linarith only [hn] + any_goals omega repeat' rw [zero_add] @[simp] diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index 75b5f6c8a09e8..c845f551fee49 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -312,9 +312,8 @@ protected theorem Commute.mul_geom_sum₂_Ico [Ring α] {x y : α} (h : Commute refine sum_congr rfl fun j j_in => ?_ rw [← pow_add] congr - rw [mem_range, Nat.lt_iff_add_one_le, add_comm] at j_in - have h' : n - m + (m - (1 + j)) = n - (1 + j) := tsub_add_tsub_cancel hmn j_in - rw [← tsub_add_eq_tsub_tsub m, h', ← tsub_add_eq_tsub_tsub] + rw [mem_range] at j_in + omega rw [this] simp_rw [pow_mul_comm y (n - m) _] simp_rw [← mul_assoc] @@ -329,10 +328,8 @@ protected theorem Commute.geom_sum₂_succ_eq {α : Type u} [Ring α] {x y : α} (h.symm.pow_right _).eq, mul_assoc, ← pow_succ'] refine sum_congr rfl fun i hi => ?_ suffices n - 1 - i + 1 = n - i by rw [this] - rcases n with - | n - · exact absurd (List.mem_range.mp hi) i.not_lt_zero - · rw [tsub_add_eq_add_tsub (Nat.le_sub_one_of_lt (List.mem_range.mp hi)), - tsub_add_cancel_of_le (Nat.succ_le_iff.mpr n.succ_pos)] + rw [Finset.mem_range] at hi + omega theorem geom_sum₂_succ_eq {α : Type u} [CommRing α] (x y : α) {n : ℕ} : ∑ i ∈ range (n + 1), x ^ i * y ^ (n - i) = diff --git a/Mathlib/Algebra/Homology/ExactSequence.lean b/Mathlib/Algebra/Homology/ExactSequence.lean index f97f59ea577c2..7b14070825673 100644 --- a/Mathlib/Algebra/Homology/ExactSequence.lean +++ b/Mathlib/Algebra/Homology/ExactSequence.lean @@ -75,7 +75,7 @@ lemma isComplex₀ (S : ComposableArrows C 0) : S.IsComplex where zero i hi := by simp +decide at hi lemma isComplex₁ (S : ComposableArrows C 1) : S.IsComplex where - zero i hi := by exfalso; omega + zero i hi := by omega variable (S) diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean b/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean index 8707a528903f7..1a1ccf798106e 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean @@ -529,8 +529,7 @@ lemma δ_ofHom {p : ℤ} (φ : F ⟶ G) : δ 0 p (Cochain.ofHom φ) = 0 := by ext simp · rw [δ_shape] - intro - exact h (by omega) + omega @[simp] lemma δ_ofHomotopy {φ₁ φ₂ : F ⟶ G} (h : Homotopy φ₁ φ₂) : diff --git a/Mathlib/Algebra/Polynomial/Module/Basic.lean b/Mathlib/Algebra/Polynomial/Module/Basic.lean index ce68ea47f409d..882036b915301 100644 --- a/Mathlib/Algebra/Polynomial/Module/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Module/Basic.lean @@ -150,9 +150,7 @@ theorem smul_single_apply (i : ℕ) (f : R[X]) (m : M) (n : ℕ) : · rw [monomial_smul_single, single_apply, coeff_monomial, ite_smul, zero_smul] by_cases h : i ≤ n · simp_rw [eq_tsub_iff_add_eq_of_le h, if_pos h] - · rw [if_neg h, ite_eq_right_iff] - intro e - exfalso + · rw [if_neg h, if_neg] omega theorem smul_apply (f : R[X]) (g : PolynomialModule R M) (n : ℕ) : diff --git a/Mathlib/Algebra/Polynomial/UnitTrinomial.lean b/Mathlib/Algebra/Polynomial/UnitTrinomial.lean index 7af5a7c7d2280..2e735b5d986a1 100644 --- a/Mathlib/Algebra/Polynomial/UnitTrinomial.lean +++ b/Mathlib/Algebra/Polynomial/UnitTrinomial.lean @@ -247,14 +247,8 @@ theorem irreducible_aux2 {k m m' n : ℕ} (hkm : k < m) (hmn : m < n) (hkm' : k · refine Or.inr ?_ rw [← trinomial_mirror hkm' hmn' u.ne_zero u.ne_zero, eq_comm, mirror_eq_iff] at hp exact hq.trans hp - · suffices m = m' by - rw [this] at hp - exact Or.inl (hq.trans hp.symm) - rw [tsub_add_eq_add_tsub hmn.le, eq_tsub_iff_add_eq_of_le, ← two_mul] at hm - · rw [tsub_add_eq_add_tsub hmn'.le, eq_tsub_iff_add_eq_of_le, ← two_mul] at hm' - · exact mul_left_cancel₀ two_ne_zero (hm.trans hm'.symm) - · exact hmn'.le.trans (Nat.le_add_right n k) - · exact hmn.le.trans (Nat.le_add_right n k) + · obtain rfl : m = m' := by omega + exact Or.inl (hq.trans hp.symm) theorem irreducible_aux3 {k m m' n : ℕ} (hkm : k < m) (hmn : m < n) (hkm' : k < m') (hmn' : m' < n) (u v w x z : Units ℤ) (hp : p = trinomial k m n (u : ℤ) v w) diff --git a/Mathlib/AlgebraicTopology/DoldKan/Faces.lean b/Mathlib/AlgebraicTopology/DoldKan/Faces.lean index 3a46ea7d170ba..9cac89a7766c3 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Faces.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Faces.lean @@ -182,10 +182,7 @@ theorem induction {Y : C} {n q : ℕ} {φ : Y ⟶ X _[n + 1]} (v : HigherFacesVa rfl -- now, we assume j ≠ a (i.e. a < j) have haj : a < j := (Ne.le_iff_lt hj₂).mp (by omega) - have ham : a ≤ m := by - by_contra h - rw [not_le, ← Nat.succ_le_iff] at h - omega + have ham : a ≤ m := by omega rw [X.δ_comp_σ_of_gt', j.pred_succ] swap · rw [Fin.lt_iff_val_lt_val] diff --git a/Mathlib/Analysis/Convex/Between.lean b/Mathlib/Analysis/Convex/Between.lean index 977702d9c6a6b..5e5d497165391 100644 --- a/Mathlib/Analysis/Convex/Between.lean +++ b/Mathlib/Analysis/Convex/Between.lean @@ -519,10 +519,7 @@ theorem sbtw_of_sbtw_of_sbtw_of_mem_affineSpan_pair [NoZeroSMulDivisors R V] have h₂₃ : i₂ ≠ i₃ := by rintro rfl simp at h₁ - have h3 : ∀ i : Fin 3, i = i₁ ∨ i = i₂ ∨ i = i₃ := by - clear h₁ h₂ h₁' h₂' - -- Porting note: Originally `decide!` - revert i₁ i₂ i₃; decide + have h3 : ∀ i : Fin 3, i = i₁ ∨ i = i₂ ∨ i = i₃ := by omega have hu : (Finset.univ : Finset (Fin 3)) = {i₁, i₂, i₃} := by clear h₁ h₂ h₁' h₂' -- Porting note: Originally `decide!` diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index c12c5f787ed51..ff1dbf10a31ac 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -802,7 +802,7 @@ theorem quadratic_root_cos_pi_div_five : push_neg intro n hn replace hn : n * 5 = 1 := by field_simp [mul_comm _ π, mul_assoc] at hn; norm_cast at hn - rcases Int.mul_eq_one_iff_eq_one_or_neg_one.mp hn with ⟨_, h⟩ | ⟨_, h⟩ <;> norm_num at h + omega suffices s * (2 * c) = s * (4 * c ^ 2 - 1) from mul_left_cancel₀ hs this calc s * (2 * c) = 2 * s * c := by rw [← mul_assoc, mul_comm 2] _ = sin (2 * θ) := by rw [sin_two_mul] diff --git a/Mathlib/CategoryTheory/ComposableArrows.lean b/Mathlib/CategoryTheory/ComposableArrows.lean index 0fcf2d10bbac5..45cd88fc81f69 100644 --- a/Mathlib/CategoryTheory/ComposableArrows.lean +++ b/Mathlib/CategoryTheory/ComposableArrows.lean @@ -143,18 +143,10 @@ lemma map_id (i : Fin 2) : map f i i (by simp) = 𝟙 _ := | 1 => rfl lemma map_comp {i j k : Fin 2} (hij : i ≤ j) (hjk : j ≤ k) : - map f i k (hij.trans hjk) = map f i j hij ≫ map f j k hjk := - match i with - | 0 => - match j with - | 0 => by rw [map_id, id_comp] - | 1 => by - obtain rfl : k = 1 := k.eq_one_of_neq_zero (by rintro rfl; simp at hjk) - rw [map_id, comp_id] - | 1 => by - obtain rfl := j.eq_one_of_neq_zero (by rintro rfl; simp at hij) - obtain rfl := k.eq_one_of_neq_zero (by rintro rfl; simp at hjk) - rw [map_id, id_comp] + map f i k (hij.trans hjk) = map f i j hij ≫ map f j k hjk := by + obtain rfl | rfl : i = j ∨ j = k := by omega + · rw [map_id, id_comp] + · rw [map_id, comp_id] end Mk₁ diff --git a/Mathlib/CategoryTheory/Preadditive/Schur.lean b/Mathlib/CategoryTheory/Preadditive/Schur.lean index 9c571b370957f..c5eb3daecf5ef 100644 --- a/Mathlib/CategoryTheory/Preadditive/Schur.lean +++ b/Mathlib/CategoryTheory/Preadditive/Schur.lean @@ -184,11 +184,8 @@ theorem finrank_hom_simple_simple_eq_zero_iff (X Y : C) [FiniteDimensional 𝕜 [FiniteDimensional 𝕜 (X ⟶ Y)] [Simple X] [Simple Y] : finrank 𝕜 (X ⟶ Y) = 0 ↔ IsEmpty (X ≅ Y) := by rw [← not_nonempty_iff, ← not_congr (finrank_hom_simple_simple_eq_one_iff 𝕜 X Y)] - refine ⟨fun h => by rw [h]; simp, fun h => ?_⟩ have := finrank_hom_simple_simple_le_one 𝕜 X Y - interval_cases finrank 𝕜 (X ⟶ Y) - · rfl - · exact False.elim (h rfl) + omega open scoped Classical diff --git a/Mathlib/Combinatorics/Enumerative/Catalan.lean b/Mathlib/Combinatorics/Enumerative/Catalan.lean index d7dd958d1d693..b69b6c29f6a33 100644 --- a/Mathlib/Combinatorics/Enumerative/Catalan.lean +++ b/Mathlib/Combinatorics/Enumerative/Catalan.lean @@ -117,7 +117,7 @@ theorem catalan_eq_centralBinom_div (n : ℕ) : catalan n = n.centralBinom / (n (Nat.centralBinom (d - i) / (d - i + 1)) : ℚ) · congr ext1 x - have m_le_d : x.val ≤ d := by apply Nat.le_of_lt_succ; apply x.2 + have m_le_d : x.val ≤ d := by omega have d_minus_x_le_d : (d - x.val) ≤ d := tsub_le_self rw [hd _ m_le_d, hd _ d_minus_x_le_d] norm_cast diff --git a/Mathlib/Combinatorics/Enumerative/DyckWord.lean b/Mathlib/Combinatorics/Enumerative/DyckWord.lean index 5924d03e54bd1..5a8e966e55c59 100644 --- a/Mathlib/Combinatorics/Enumerative/DyckWord.lean +++ b/Mathlib/Combinatorics/Enumerative/DyckWord.lean @@ -205,8 +205,7 @@ def denest (hn : p.IsNested) : DyckWord where 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 lb : 0 < min (1 + i) (p.toList.length - 1) := by omega have eq := hn.2 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, diff --git a/Mathlib/Combinatorics/SimpleGraph/Walk.lean b/Mathlib/Combinatorics/SimpleGraph/Walk.lean index c7330ac16a9ad..ae8224b0ec59d 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Walk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Walk.lean @@ -1115,9 +1115,8 @@ theorem mem_support_iff_exists_getVert {u v w : V} {p : G.Walk v w} : exact hn.1.symm · right have hnp : ¬ p.Nil := by - rw [@nil_iff_length_eq] - have : 1 ≤ p.length := by omega - exact Nat.not_eq_zero_of_lt this + rw [nil_iff_length_eq] + omega rw [← support_tail_of_not_nil _ hnp] rw [mem_support_iff_exists_getVert] use n - 1 diff --git a/Mathlib/Data/FP/Basic.lean b/Mathlib/Data/FP/Basic.lean index 492cf4ffcd604..38182b202701d 100644 --- a/Mathlib/Data/FP/Basic.lean +++ b/Mathlib/Data/FP/Basic.lean @@ -5,7 +5,11 @@ Authors: Mario Carneiro -/ import Mathlib.Data.Semiquot import Mathlib.Data.Nat.Size -import Mathlib.Tactic.Ring.RingNF +import Batteries.Data.Rat.Basic +import Mathlib.Data.PNat.Defs +import Mathlib.Data.Rat.Init +import Mathlib.Algebra.Ring.Int.Defs +import Mathlib.Algebra.Order.Group.Unbundled.Basic /-! # Implementation of floating-point numbers (experimental). @@ -82,9 +86,7 @@ theorem Float.Zero.valid : ValidFinite emin 0 := rw [← Int.ofNat_le] at this rw [← sub_nonneg] at * simp only [emin, emax] at * - ring_nf - rw [mul_comm] - assumption + omega le_trans C.precMax (Nat.le_mul_of_pos_left _ Nat.zero_lt_two), by (rw [max_eq_right]; simp [sub_eq_add_neg, Int.ofNat_zero_le])⟩ diff --git a/Mathlib/Data/Int/Defs.lean b/Mathlib/Data/Int/Defs.lean index eebbbbab270c4..3addc51ed7381 100644 --- a/Mathlib/Data/Int/Defs.lean +++ b/Mathlib/Data/Int/Defs.lean @@ -370,21 +370,11 @@ lemma ediv_of_neg_of_pos {a b : ℤ} (Ha : a < 0) (Hb : 0 < b) : ediv a b = -((- lemma add_emod_eq_add_mod_right {m n k : ℤ} (i : ℤ) (H : m % n = k % n) : (m + i) % n = (k + i) % n := by rw [← emod_add_emod, ← emod_add_emod k, H] -@[simp] lemma neg_emod_two (i : ℤ) : -i % 2 = i % 2 := by - apply Int.emod_eq_emod_iff_emod_sub_eq_zero.mpr - convert Int.mul_emod_right 2 (-i) using 2 - rw [Int.two_mul, Int.sub_eq_add_neg] +@[simp] lemma neg_emod_two (i : ℤ) : -i % 2 = i % 2 := by omega /-! ### properties of `/` and `%` -/ -lemma emod_two_eq_zero_or_one (n : ℤ) : n % 2 = 0 ∨ n % 2 = 1 := - have h : n % 2 < 2 := by omega - have h₁ : 0 ≤ n % 2 := Int.emod_nonneg _ (by decide) - match n % 2, h, h₁ with - | (0 : ℕ), _ ,_ => Or.inl rfl - | (1 : ℕ), _ ,_ => Or.inr rfl - | (k + 2 : ℕ), h₁, _ => by omega - | -[a+1], _, h₁ => by cases h₁ +lemma emod_two_eq_zero_or_one (n : ℤ) : n % 2 = 0 ∨ n % 2 = 1 := by omega /-! ### dvd -/ @@ -542,7 +532,7 @@ lemma sign_add_eq_of_sign_eq : ∀ {m n : ℤ}, m.sign = n.sign → (m + n).sign have : (1 : ℤ) ≠ -1 := by decide rintro ((_ | m) | m) ((_ | n) | n) <;> simp [this, this.symm, Int.negSucc_add_negSucc] rw [Int.sign_eq_one_iff_pos] - apply Int.add_pos <;> omega + omega /-! ### toNat -/ diff --git a/Mathlib/Data/List/NodupEquivFin.lean b/Mathlib/Data/List/NodupEquivFin.lean index b53d6dbaf8777..9b381e748861c 100644 --- a/Mathlib/Data/List/NodupEquivFin.lean +++ b/Mathlib/Data/List/NodupEquivFin.lean @@ -185,8 +185,7 @@ theorem sublist_iff_exists_fin_orderEmbedding_get_eq {l l' : List α} : dsimp only split_ifs with hi hj hj · rwa [Fin.val_fin_lt, f.lt_iff_lt] - · have := (f ⟨i, hi⟩).is_lt - omega + · omega · exact absurd (h.trans hj) hi · simpa using h · intro i diff --git a/Mathlib/Data/List/Rotate.lean b/Mathlib/Data/List/Rotate.lean index 55830f40c29b9..1e658d0fb7491 100644 --- a/Mathlib/Data/List/Rotate.lean +++ b/Mathlib/Data/List/Rotate.lean @@ -206,9 +206,7 @@ theorem getElem?_rotate {l : List α} {n m : ℕ} (hml : m < l.length) : rw [Nat.sub_lt_iff_lt_add' hm'] exact Nat.add_lt_add hlt hml conv_rhs => rw [Nat.add_comm m, ← mod_add_mod, mod_eq_sub_mod hm', mod_eq_of_lt this] - rw [← Nat.add_right_inj, ← Nat.add_sub_assoc, Nat.add_sub_sub_cancel, Nat.add_sub_cancel', - Nat.add_comm] - exacts [hm', hlt.le, hm] + omega · rwa [Nat.sub_lt_iff_lt_add hm, length_drop, Nat.sub_add_cancel hlt.le] theorem getElem_rotate (l : List α) (n : ℕ) (k : Nat) (h : k < (l.rotate n).length) : diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 05f6a45500d6c..e1d2bb4909d02 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -245,8 +245,7 @@ lemma add_eq_min_iff : m + n = min m n ↔ m = 0 ∧ n = 0 := by omega lemma add_pos_iff_pos_or_pos : 0 < m + n ↔ 0 < m ∨ 0 < n := by omega -lemma add_eq_one_iff : m + n = 1 ↔ m = 0 ∧ n = 1 ∨ m = 1 ∧ n = 0 := by - cases n <;> simp [← Nat.add_assoc, succ_inj'] +lemma add_eq_one_iff : m + n = 1 ↔ m = 0 ∧ n = 1 ∨ m = 1 ∧ n = 0 := by omega lemma add_eq_two_iff : m + n = 2 ↔ m = 0 ∧ n = 2 ∨ m = 1 ∧ n = 1 ∨ m = 2 ∧ n = 0 := by omega @@ -255,27 +254,13 @@ lemma add_eq_three_iff : m + n = 3 ↔ m = 0 ∧ n = 3 ∨ m = 1 ∧ n = 2 ∨ m = 2 ∧ n = 1 ∨ m = 3 ∧ n = 0 := by omega -lemma le_add_one_iff : m ≤ n + 1 ↔ m ≤ n ∨ m = n + 1 := by - rw [Nat.le_iff_lt_or_eq, Nat.lt_add_one_iff] +lemma le_add_one_iff : m ≤ n + 1 ↔ m ≤ n ∨ m = n + 1 := by omega -lemma le_and_le_add_one_iff : n ≤ m ∧ m ≤ n + 1 ↔ m = n ∨ m = n + 1 := by - rw [le_add_one_iff, and_or_left, ← Nat.le_antisymm_iff, eq_comm, and_iff_right_of_imp] - rintro rfl - exact n.le_succ +lemma le_and_le_add_one_iff : n ≤ m ∧ m ≤ n + 1 ↔ m = n ∨ m = n + 1 := by omega -lemma add_succ_lt_add (hab : a < b) (hcd : c < d) : a + c + 1 < b + d := by - rw [Nat.add_assoc]; exact Nat.add_lt_add_of_lt_of_le hab (Nat.succ_le_iff.2 hcd) +lemma add_succ_lt_add (hab : a < b) (hcd : c < d) : a + c + 1 < b + d := by omega -theorem le_or_le_of_add_eq_add_pred (h : a + c = b + d - 1) : b ≤ a ∨ d ≤ c := by - rcases le_or_lt b a with h' | h' <;> [left; right] - · exact h' - · replace h' := Nat.add_lt_add_right h' c - rw [h] at h' - rcases d.eq_zero_or_pos with hn | hn - · rw [hn] - exact zero_le c - rw [d.add_sub_assoc (Nat.succ_le_of_lt hn), Nat.add_lt_add_iff_left] at h' - exact Nat.le_of_pred_lt h' +theorem le_or_le_of_add_eq_add_pred (h : a + c = b + d - 1) : b ≤ a ∨ d ≤ c := by omega /-! ### `sub` -/ @@ -285,9 +270,9 @@ attribute [simp] Nat.sub_eq_zero_of_le Nat.sub_le_iff_le_add Nat.add_sub_cancel_ /-- A version of `Nat.sub_succ` in the form `_ - 1` instead of `Nat.pred _`. -/ lemma sub_succ' (m n : ℕ) : m - n.succ = m - n - 1 := rfl -protected lemma sub_eq_of_eq_add' (h : a = b + c) : a - b = c := by rw [h, Nat.add_sub_cancel_left] -protected lemma eq_sub_of_add_eq (h : c + b = a) : c = a - b := (Nat.sub_eq_of_eq_add h.symm).symm -protected lemma eq_sub_of_add_eq' (h : b + c = a) : c = a - b := (Nat.sub_eq_of_eq_add' h.symm).symm +protected lemma sub_eq_of_eq_add' (h : a = b + c) : a - b = c := by omega +protected lemma eq_sub_of_add_eq (h : c + b = a) : c = a - b := by omega +protected lemma eq_sub_of_add_eq' (h : b + c = a) : c = a - b := by omega protected lemma lt_sub_iff_add_lt : a < c - b ↔ a + b < c := ⟨add_lt_of_lt_sub, lt_sub_of_add_lt⟩ protected lemma lt_sub_iff_add_lt' : a < c - b ↔ b + a < c := by omega diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index c2cdfeb2c6e50..5560ced0e6857 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -83,12 +83,7 @@ theorem Squarefree.ext_iff {n m : ℕ} (hn : Squarefree n) (hm : Squarefree m) : hp.dvd_iff_one_le_factorization hm.ne_zero, not_le, lt_one_iff] at h₁ have h₂ := hn.natFactorization_le_one p have h₃ := hm.natFactorization_le_one p - rw [Nat.le_add_one_iff, Nat.le_zero] at h₂ h₃ - cases' h₂ with h₂ h₂ - · rwa [h₂, eq_comm, ← h₁] - · rw [h₂, h₃.resolve_left] - rw [← h₁, h₂] - simp only [Nat.one_ne_zero, not_false_iff, reduceCtorEq] + omega rw [factorization_eq_zero_of_non_prime _ hp, factorization_eq_zero_of_non_prime _ hp] theorem squarefree_pow_iff {n k : ℕ} (hn : n ≠ 1) (hk : k ≠ 0) : diff --git a/Mathlib/Data/Ordmap/Ordset.lean b/Mathlib/Data/Ordmap/Ordset.lean index 5b497cb5df020..eee6810eabc6b 100644 --- a/Mathlib/Data/Ordmap/Ordset.lean +++ b/Mathlib/Data/Ordmap/Ordset.lean @@ -1147,10 +1147,7 @@ theorem Valid'.balance'_aux {l} {x : α} {r o₁ o₂} (hl : Valid' o₁ l x) (h theorem Valid'.balance'_lemma {α l l' r r'} (H1 : BalancedSz l' r') (H2 : Nat.dist (@size α l) l' ≤ 1 ∧ size r = r' ∨ Nat.dist (size r) r' ≤ 1 ∧ size l = l') : 2 * @size α r ≤ 9 * size l + 5 ∨ size r ≤ 3 := by - suffices @size α r ≤ 3 * (size l + 1) by - rcases Nat.eq_zero_or_pos (size l) with l0 | l0 - · apply Or.inr; rwa [l0] at this - change 1 ≤ _ at l0; apply Or.inl; omega + suffices @size α r ≤ 3 * (size l + 1) by omega rcases H2 with (⟨hl, rfl⟩ | ⟨hr, rfl⟩) <;> rcases H1 with (h | ⟨_, h₂⟩) · exact le_trans (Nat.le_add_left _ _) (le_trans h (Nat.le_add_left _ _)) · exact diff --git a/Mathlib/Data/Real/Archimedean.lean b/Mathlib/Data/Real/Archimedean.lean index a42f86a993916..1fc5a7ad21d35 100644 --- a/Mathlib/Data/Real/Archimedean.lean +++ b/Mathlib/Data/Real/Archimedean.lean @@ -373,8 +373,7 @@ lemma exists_natCast_add_one_lt_pow_of_one_lt (ha : 1 < a) : ∃ m : ℕ, (m + 1 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 + omega use 2 * k ^ 2 calc ((2 * k ^ 2 : ℕ) + 1 : ℝ) ≤ 2 ^ (2 * k) := mod_cast Nat.two_mul_sq_add_one_le_two_pow_two_mul _ diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 37dffbc1303eb..05692bc1135dc 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -1268,20 +1268,8 @@ variable {n a : ℕ} theorem valMinAbs_natAbs_eq_min {n : ℕ} [hpos : NeZero n] (a : ZMod n) : a.valMinAbs.natAbs = min a.val (n - a.val) := by rw [valMinAbs_def_pos] - split_ifs with h - · rw [Int.natAbs_ofNat] - symm - apply - min_eq_left (le_trans h (le_trans (Nat.half_le_of_sub_le_half _) (Nat.sub_le_sub_left h n))) - rw [Nat.sub_sub_self (Nat.div_le_self _ _)] - · rw [← Int.natAbs_neg, neg_sub, ← Nat.cast_sub a.val_le] - symm - apply - min_eq_right - (le_trans (le_trans (Nat.sub_le_sub_left (lt_of_not_ge h) n) (Nat.le_half_of_half_lt_sub _)) - (le_of_not_ge h)) - rw [Nat.sub_sub_self (Nat.div_lt_self (lt_of_le_of_ne' (Nat.zero_le _) hpos.1) one_lt_two)] - apply Nat.lt_succ_self + have := a.val_lt + omega theorem valMinAbs_natCast_of_le_half (ha : a ≤ n / 2) : (a : ZMod n).valMinAbs = a := by cases n diff --git a/Mathlib/Geometry/Euclidean/MongePoint.lean b/Mathlib/Geometry/Euclidean/MongePoint.lean index 9800925e6fb48..038fc99fd056f 100644 --- a/Mathlib/Geometry/Euclidean/MongePoint.lean +++ b/Mathlib/Geometry/Euclidean/MongePoint.lean @@ -468,12 +468,8 @@ theorem eq_orthocenter_of_forall_mem_altitude {t : Triangle ℝ P} {i₁ i₂ : rw [orthocenter_eq_mongePoint] have ha : ∀ i, i₃ ≠ i → p ∈ t.mongePlane i₃ i := by intro i hi - have hi₁₂ : i₁ = i ∨ i₂ = i := by - clear h₁ h₂ - decide +revert - cases' hi₁₂ with hi₁₂ hi₁₂ - · exact hi₁₂ ▸ h₂ - · exact hi₁₂ ▸ h₁ + obtain rfl | rfl : i₁ = i ∨ i₂ = i := by omega + all_goals assumption exact eq_mongePoint_of_forall_mem_mongePlane ha /-- The distance from the orthocenter to the reflection of the diff --git a/Mathlib/GroupTheory/Coxeter/Length.lean b/Mathlib/GroupTheory/Coxeter/Length.lean index 43e381a2836af..fbc5887d029a5 100644 --- a/Mathlib/GroupTheory/Coxeter/Length.lean +++ b/Mathlib/GroupTheory/Coxeter/Length.lean @@ -246,10 +246,7 @@ theorem not_isReduced_alternatingWord (i i' : B) {m : ℕ} (hM : M i i' ≠ 0) ( omega have : M i i' + 1 ≤ M i i' * 2 := by linarith [Nat.one_le_iff_ne_zero.mpr hM] rw [cs.prod_alternatingWord_eq_prod_alternatingWord_sub i i' _ this] - have : M i i' * 2 - (M i i' + 1) = M i i' - 1 := by - apply (Nat.sub_eq_iff_eq_add' this).mpr - rw [add_assoc, add_comm 1, Nat.sub_add_cancel (Nat.one_le_iff_ne_zero.mpr hM)] - exact mul_two _ + have : M i i' * 2 - (M i i' + 1) = M i i' - 1 := by omega rw [this] calc ℓ (π (alternatingWord i' i (M i i' - 1))) @@ -306,8 +303,7 @@ theorem isLeftDescent_iff {w : W} {i : B} : constructor · intro _ exact (cs.length_simple_mul w i).resolve_left (by omega) - · intro _ - omega + · omega theorem not_isLeftDescent_iff {w : W} {i : B} : ¬cs.IsLeftDescent w i ↔ ℓ (s i * w) = ℓ w + 1 := by @@ -315,8 +311,7 @@ theorem not_isLeftDescent_iff {w : W} {i : B} : constructor · intro _ exact (cs.length_simple_mul w i).resolve_right (by omega) - · intro _ - omega + · omega theorem isRightDescent_iff {w : W} {i : B} : cs.IsRightDescent w i ↔ ℓ (w * s i) + 1 = ℓ w := by @@ -324,8 +319,7 @@ theorem isRightDescent_iff {w : W} {i : B} : constructor · intro _ exact (cs.length_mul_simple w i).resolve_left (by omega) - · intro _ - omega + · omega theorem not_isRightDescent_iff {w : W} {i : B} : ¬cs.IsRightDescent w i ↔ ℓ (w * s i) = ℓ w + 1 := by @@ -333,8 +327,7 @@ theorem not_isRightDescent_iff {w : W} {i : B} : constructor · intro _ exact (cs.length_mul_simple w i).resolve_right (by omega) - · intro _ - omega + · omega theorem isLeftDescent_iff_not_isLeftDescent_mul {w : W} {i : B} : cs.IsLeftDescent w i ↔ ¬cs.IsLeftDescent (s i * w) i := by diff --git a/Mathlib/GroupTheory/FreeGroup/Basic.lean b/Mathlib/GroupTheory/FreeGroup/Basic.lean index a2174bda20dfc..e0ee2f9a2a72d 100644 --- a/Mathlib/GroupTheory/FreeGroup/Basic.lean +++ b/Mathlib/GroupTheory/FreeGroup/Basic.lean @@ -336,18 +336,8 @@ theorem sizeof_of_step : ∀ {L₁ L₂ : List (α × Bool)}, | _, _, @Step.not _ L1 L2 x b => by induction L1 with | nil => - -- dsimp [sizeOf] dsimp - simp only [Bool.sizeOf_eq_one] - - have H : - 1 + (1 + 1) + (1 + (1 + 1) + sizeOf L2) = - sizeOf L2 + (1 + ((1 + 1) + (1 + 1) + 1)) := by - ac_rfl - rw [H] - apply Nat.lt_add_of_pos_right - apply Nat.lt_add_right - apply Nat.zero_lt_one + omega | cons hd tl ih => dsimp exact Nat.add_lt_add_left ih _ diff --git a/Mathlib/GroupTheory/FreeGroup/Reduce.lean b/Mathlib/GroupTheory/FreeGroup/Reduce.lean index 3017c17281dd2..de90d1c5721c6 100644 --- a/Mathlib/GroupTheory/FreeGroup/Reduce.lean +++ b/Mathlib/GroupTheory/FreeGroup/Reduce.lean @@ -95,7 +95,6 @@ theorem reduce.not {p : Prop} : have := congr_arg List.length h simp? [List.length] at this says simp only [List.length, zero_add, List.length_append] at this - rw [add_comm, add_assoc, add_assoc, add_comm, <-add_assoc] at this omega | cons hd tail => cases' hd with y c diff --git a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean index 7dd9feaf90d4d..df3d8243d1cdb 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean @@ -265,7 +265,7 @@ theorem isConj_swap_mul_swap_of_cycleType_two {g : Perm (Fin 5)} (ha : g ∈ alt rw [← Multiset.eq_replicate_card] at h2 rw [← sum_cycleType, h2, Multiset.sum_replicate, smul_eq_mul] at h have h : Multiset.card g.cycleType ≤ 3 := - le_of_mul_le_mul_right (le_trans h (by simp only [card_fin]; ring_nf; decide)) (by simp) + le_of_mul_le_mul_right (le_trans h (by norm_num only [card_fin])) (by simp) rw [mem_alternatingGroup, sign_of_cycleType, h2] at ha norm_num at ha rw [pow_add, pow_mul, Int.units_pow_two, one_mul, neg_one_pow_eq_one_iff_even] at ha diff --git a/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean b/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean index c003e3f2279c9..6cebe6c4de152 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean @@ -325,14 +325,12 @@ lemma finrank_orthogonal (hB : B.Nondegenerate) (hB₀ : B.IsRefl) (W : Submodul finrank K (B.orthogonal W) = finrank K V - finrank K W := by have := finrank_add_finrank_orthogonal hB₀ (W := W) rw [B.orthogonal_top_eq_bot hB hB₀, inf_bot_eq, finrank_bot, add_zero] at this - have : finrank K W ≤ finrank K V := finrank_le W omega lemma orthogonal_orthogonal (hB : B.Nondegenerate) (hB₀ : B.IsRefl) (W : Submodule K V) : B.orthogonal (B.orthogonal W) = W := by apply (eq_of_le_of_finrank_le (LinearMap.BilinForm.le_orthogonal_orthogonal hB₀) _).symm simp only [finrank_orthogonal hB hB₀] - have : finrank K W ≤ finrank K V := finrank_le W omega variable {W : Submodule K V} diff --git a/Mathlib/LinearAlgebra/Reflection.lean b/Mathlib/LinearAlgebra/Reflection.lean index fdeda916180ba..7201b2603f574 100644 --- a/Mathlib/LinearAlgebra/Reflection.lean +++ b/Mathlib/LinearAlgebra/Reflection.lean @@ -224,14 +224,8 @@ lemma reflection_mul_reflection_zpow_apply (m : ℤ) (z : M) have ht' : t = g x * f y - 2 := by rwa [mul_comm (g x)] rw [zpow_neg, ← inv_zpow, mul_inv_rev, reflection_inv, reflection_inv, zpow_natCast, reflection_mul_reflection_pow_apply hg hf m z t ht', add_right_comm z] - have aux {a b : ℤ} (hab : a + b = -3) : a / 2 = -(b / 2) - 2 := by - rw [← mul_right_inj' (by norm_num : (2 : ℤ) ≠ 0), mul_sub, mul_neg, - eq_sub_of_add_eq (Int.ediv_add_emod _ _), eq_sub_of_add_eq (Int.ediv_add_emod _ _)] - omega - rw [aux (by omega : (-m - 3) + m = (-3 : ℤ)), - aux (by omega : (-m - 2) + (m - 1) = (-3 : ℤ)), - aux (by omega : (-m - 1) + (m - 2) = (-3 : ℤ)), - aux (by omega : -m + (m - 3) = (-3 : ℤ))] + have aux (a b : ℤ) (hab : a + b = -3 := by omega) : a / 2 = -(b / 2) - 2 := by omega + rw [aux (-m - 3) m, aux (-m - 2) (m - 1), aux (-m - 1) (m - 2), aux (-m) (m - 3)] simp only [S_neg_sub_two, Polynomial.eval_neg] ring_nf diff --git a/Mathlib/ModelTheory/Semantics.lean b/Mathlib/ModelTheory/Semantics.lean index a9b3f5e826706..558f312224a96 100644 --- a/Mathlib/ModelTheory/Semantics.lean +++ b/Mathlib/ModelTheory/Semantics.lean @@ -383,16 +383,8 @@ theorem realize_liftAt {n n' m : ℕ} {φ : L.BoundedFormula α n} {v : α → M simp only [mapTermRel, Realize, realize_castLE_of_eq h, ih3 (hmn.trans k.succ.le_succ)] refine forall_congr' fun x => iff_eq_eq.mpr (congr rfl (funext (Fin.lastCases ?_ fun i => ?_))) · simp only [Function.comp_apply, val_last, snoc_last] - by_cases h : k < m - · rw [if_pos h] - refine (congr rfl (Fin.ext ?_)).trans (snoc_last _ _) - simp only [coe_cast, coe_castAdd, val_last, self_eq_add_right] - refine le_antisymm - (le_of_add_le_add_left ((hmn.trans (Nat.succ_le_of_lt h)).trans ?_)) n'.zero_le - rw [add_zero] - · rw [if_neg h] - refine (congr rfl (Fin.ext ?_)).trans (snoc_last _ _) - simp + refine (congr rfl (Fin.ext ?_)).trans (snoc_last _ _) + split_ifs <;> dsimp; omega · simp only [Function.comp_apply, Fin.snoc_castSucc] refine (congr rfl (Fin.ext ?_)).trans (snoc_castSucc _ _ _) simp only [coe_castSucc, coe_cast] diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean index 360acae7e6d20..c99421dacd8f4 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean @@ -290,8 +290,6 @@ theorem FiniteField.isSquare_neg_one_iff : IsSquare (-1 : F) ↔ Fintype.card F · have h₁ := FiniteField.odd_card_of_char_ne_two hF rw [← quadraticChar_one_iff_isSquare (neg_ne_zero.mpr (one_ne_zero' F)), quadraticChar_neg_one hF, χ₄_nat_eq_if_mod_four, h₁] - simp only [Nat.one_ne_zero, if_false, ite_eq_left_iff, (by omega : (-1 : ℤ) ≠ 1), imp_false, - not_not, Ne, reduceCtorEq] - exact ⟨fun h ↦ ne_of_eq_of_ne h (by omega), (Nat.odd_mod_four_iff.mp h₁).resolve_right⟩ + omega end SpecialValues diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean index e29cef14672bb..d09b570d395c7 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean @@ -70,21 +70,13 @@ namespace ZMod 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₁ - have h₂ := mod_lt p (by norm_num : 0 < 8) - revert h₂ h₁ - generalize p % 8 = m; clear! p - intros; interval_cases m <;> simp_all -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11043): was `decide!` + omega /-- `-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 (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₁ - have h₂ := mod_lt p (by norm_num : 0 < 8) - revert h₂ h₁ - generalize p % 8 = m; clear! p - intros; interval_cases m <;> simp_all -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11043): was `decide!` + omega end ZMod diff --git a/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean b/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean index 07c80e89edce7..4734feed96601 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean @@ -346,13 +346,7 @@ lemma add_eq_min {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) have h3 := min_le_padicValRat_add (p := p) (ne_of_eq_of_ne (add_neg_cancel_right r q) hr) rw [add_neg_cancel_right, padicValRat.neg] at h2 h3 rw [add_comm] at h3 - refine le_antisymm (le_min ?_ ?_) h1 - · contrapose! h2 - rw [min_eq_right h2.le] at h3 - exact lt_min h2 (lt_of_le_of_ne h3 hval) - · contrapose! h3 - rw [min_eq_right h3.le] at h2 - exact lt_min h3 (lt_of_le_of_ne h2 hval.symm) + omega lemma add_eq_of_lt {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q < padicValRat p r) : diff --git a/Mathlib/NumberTheory/Pell.lean b/Mathlib/NumberTheory/Pell.lean index 583f3cafeda19..b69923c045a69 100644 --- a/Mathlib/NumberTheory/Pell.lean +++ b/Mathlib/NumberTheory/Pell.lean @@ -213,7 +213,7 @@ theorem d_nonsquare_of_one_lt_x {a : Solution₁ d} (ha : 1 < a.x) : ¬IsSquare have hp := a.prop rintro ⟨b, rfl⟩ simp_rw [← sq, ← mul_pow, sq_sub_sq, Int.mul_eq_one_iff_eq_one_or_neg_one] at hp - rcases hp with (⟨hp₁, hp₂⟩ | ⟨hp₁, hp₂⟩) <;> omega + omega /-- A solution with `x = 1` is trivial. -/ theorem eq_one_of_x_eq_one (h₀ : d ≠ 0) {a : Solution₁ d} (ha : a.x = 1) : a = 1 := by diff --git a/Mathlib/NumberTheory/PellMatiyasevic.lean b/Mathlib/NumberTheory/PellMatiyasevic.lean index 1859b50c01808..b3c0223113d11 100644 --- a/Mathlib/NumberTheory/PellMatiyasevic.lean +++ b/Mathlib/NumberTheory/PellMatiyasevic.lean @@ -710,10 +710,7 @@ theorem eq_of_xn_modEq' {i j n} (ipos : 0 < i) (hin : i ≤ n) (j4n : j ≤ 4 * (h.symm.trans <| by let t := xn_modEq_x4n_sub a1 j42n rwa [tsub_tsub_cancel_of_le j4n] at t) - fun _ n1 => - ⟨fun i0 => absurd i0 (_root_.ne_of_gt ipos), fun i2 => by - rw [n1, i2] at hin - exact absurd hin (by decide)⟩ + (by omega) theorem modEq_of_xn_modEq {i j n} (ipos : 0 < i) (hin : i ≤ n) (h : xn a1 j ≡ xn a1 i [MOD xn a1 n]) : @@ -890,8 +887,7 @@ theorem eq_pow_of_pell {m n k} : exact x_sub_y_dvd_pow a1 n k have ta : 2 * a * n = t + (n * n + 1) := by zify - rw [te] - ring_nf + omega have zp : a * a - ((w + 1) * (w + 1) - 1) * (w * z) * (w * z) = 1 := ze ▸ pell_eq w1 w exact ⟨w, a, t, z, a1, tm, ta, Nat.cast_lt.1 nt, nw, kw, zp⟩ · rintro (⟨rfl, rfl⟩ | ⟨hk0, ⟨rfl, rfl⟩ | ⟨hn0, w, a, t, z, a1, tm, ta, mt, nw, kw, zp⟩⟩) diff --git a/Mathlib/NumberTheory/PythagoreanTriples.lean b/Mathlib/NumberTheory/PythagoreanTriples.lean index d4c366db72320..4d969268df6d1 100644 --- a/Mathlib/NumberTheory/PythagoreanTriples.lean +++ b/Mathlib/NumberTheory/PythagoreanTriples.lean @@ -194,9 +194,7 @@ theorem normalize : PythagoreanTriple (x / Int.gcd x y) (y / Int.gcd x y) (z / I theorem isClassified_of_isPrimitiveClassified (hp : h.IsPrimitiveClassified) : h.IsClassified := by obtain ⟨m, n, H⟩ := hp use 1, m, n - rcases H with ⟨t, co, _⟩ - rw [one_mul, one_mul] - exact ⟨t, co⟩ + omega theorem isClassified_of_normalize_isPrimitiveClassified (hc : h.normalize.IsPrimitiveClassified) : h.IsClassified := by @@ -531,9 +529,7 @@ theorem isPrimitiveClassified_of_coprime_of_odd_of_pos (hc : Int.gcd x y = 1) (h · apply (mul_lt_mul_right (by norm_num : 0 < (2 : ℤ))).mp rw [Int.ediv_mul_cancel h1.1, zero_mul] exact hm2n2 - rw [h2.1, h1.2.2.1] at hyo - revert hyo - norm_num + norm_num [h2.1, h1.2.2.1] at hyo theorem isPrimitiveClassified_of_coprime_of_pos (hc : Int.gcd x y = 1) (hzpos : 0 < z) : h.IsPrimitiveClassified := by diff --git a/Mathlib/NumberTheory/SumTwoSquares.lean b/Mathlib/NumberTheory/SumTwoSquares.lean index 4c8c1a2799082..ae1099ea7d72e 100644 --- a/Mathlib/NumberTheory/SumTwoSquares.lean +++ b/Mathlib/NumberTheory/SumTwoSquares.lean @@ -151,7 +151,7 @@ theorem ZMod.isSquare_neg_one_of_eq_sq_add_sq_of_isCoprime {n x y : ℤ} (h : n (hc : IsCoprime x y) : IsSquare (-1 : ZMod n.natAbs) := by obtain ⟨u, v, huv⟩ : IsCoprime x n := by have hc2 : IsCoprime (x ^ 2) (y ^ 2) := hc.pow - rw [show y ^ 2 = n + -1 * x ^ 2 by rw [h]; ring] at hc2 + rw [show y ^ 2 = n + -1 * x ^ 2 by omega] at hc2 exact (IsCoprime.pow_left_iff zero_lt_two).mp hc2.of_add_mul_right_right have H : u * y * (u * y) - -1 = n * (-v ^ 2 * n + u ^ 2 + 2 * v) := by linear_combination -u ^ 2 * h + (n * v - u * x - 1) * huv diff --git a/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean b/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean index 48892d46d0103..3703485af2737 100644 --- a/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean +++ b/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean @@ -59,7 +59,7 @@ theorem eventually_residual_liouville : ∀ᶠ x in residual ℝ, Liouville x := · rintro _ ⟨r, rfl⟩ simp only [mem_iInter, mem_iUnion] refine fun n => ⟨r.num * 2, r.den * 2, ?_, ?_⟩ - · have := Int.ofNat_le.2 r.pos; rw [Int.ofNat_one] at this; omega + · have := r.pos; omega · convert @mem_ball_self ℝ _ (r : ℝ) _ _ · push_cast; norm_cast; simp [Rat.divInt_mul_right (two_ne_zero), Rat.mkRat_self] · refine one_div_pos.2 (pow_pos (Int.cast_pos.2 ?_) _) diff --git a/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean b/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean index 586accb6d6d3c..46d933b28d37c 100644 --- a/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean +++ b/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean @@ -39,12 +39,7 @@ theorem mod_four_eq_three_of_nat_prime_of_prime (p : ℕ) [hp : Fact p.Prime] exact absurd this (by decide))) fun hp1 => by_contradiction fun hp3 : p % 4 ≠ 3 => by - have hp41 : p % 4 = 1 := by - rw [← Nat.mod_mul_left_mod p 2 2, show 2 * 2 = 4 from rfl] at hp1 - have := Nat.mod_lt p (show 0 < 4 by decide) - revert this hp3 hp1 - generalize p % 4 = m - intros; interval_cases m <;> simp_all -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11043): was `decide!` + have hp41 : p % 4 = 1 := by omega let ⟨k, hk⟩ := (ZMod.exists_sq_eq_neg_one_iff (p := p)).2 <| by rw [hp41]; decide obtain ⟨k, k_lt_p, rfl⟩ : ∃ (k' : ℕ) (_ : k' < p), (k' : ZMod p) = k := by exact ⟨k.val, k.val_lt, ZMod.natCast_zmod_val k⟩ diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index e385b759958c2..3b0362616ff38 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -256,8 +256,7 @@ def append (p q : RelSeries r) (connect : r p.last q.head) : RelSeries r where Nat.add_sub_cancel' <| le_of_lt (show p.length < i.1 from hi), add_comm] rfl rw [hx, Fin.append_right, hy, Fin.append_right] - convert q.step ⟨i - (p.length + 1), Nat.sub_lt_left_of_lt_add hi <| - by convert i.2 using 1; exact Nat.add_right_comm ..⟩ + convert q.step ⟨i - (p.length + 1), Nat.sub_lt_left_of_lt_add hi <| by omega⟩ rw [Fin.succ_mk, Nat.sub_eq_iff_eq_add (le_of_lt hi : p.length ≤ i), Nat.add_assoc _ 1, add_comm 1, Nat.sub_add_cancel] exact hi @@ -366,7 +365,7 @@ def insertNth (p : RelSeries r) (i : Fin p.length) (a : α) · change Fin.insertNth _ _ _ _ = _ rw [Fin.insertNth_apply_above] swap - · change i.1 + 1 < m.1 + 1; rw [hm]; exact lt_add_one _ + · change i.1 + 1 < m.1 + 1; omega simp only [Fin.pred_succ, eq_rec_constant] congr; ext; exact hm.symm @@ -384,8 +383,7 @@ def reverse (p : RelSeries r) : RelSeries (fun (a b : α) ↦ r b a) where convert p.step ⟨p.length - (i.1 + 1), Nat.sub_lt_self (by omega) hi⟩ · ext; simp · ext - simp only [Fin.val_rev, Fin.coe_castSucc, Nat.succ_sub_succ_eq_sub, Fin.val_succ] - rw [Nat.sub_eq_iff_eq_add, add_assoc, add_comm 1 i.1, Nat.sub_add_cancel] <;> + simp only [Fin.val_rev, Fin.coe_castSucc, Fin.val_succ] omega @[simp] lemma reverse_apply (p : RelSeries r) (i : Fin (p.length + 1)) : diff --git a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean index 99cf15144c642..0ad41dd02719a 100644 --- a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean @@ -363,7 +363,7 @@ lemma exists_eval_ne_zero_of_totalDegree_le_card_aux {N : ℕ} {F : MvPolynomial obtain hFn | hFn := ne_or_eq ((finSuccEquiv R N F).coeff n) 0 · exact hF.exists_eval_ne_zero_of_coeff_finSuccEquiv_ne_zero_aux hFn have hin : i < n := hin.lt_or_eq.elim id <| by aesop - obtain ⟨j, hj⟩ : ∃ j, i + (j + 1) = n := (Nat.exists_eq_add_of_lt hin).imp <| by intros; omega + obtain ⟨j, hj⟩ : ∃ j, i + (j + 1) = n := (Nat.exists_eq_add_of_lt hin).imp <| by omega obtain ⟨r, hr⟩ : ∃ r, (eval r) (Polynomial.coeff ((finSuccEquiv R N) F) i) ≠ 0 := IH (hF.finSuccEquiv_coeff_isHomogeneous _ _ hj) hi (.trans (by norm_cast; omega) hnR) set φ : R[X] := Polynomial.map (eval r) (finSuccEquiv _ _ F) with hφ diff --git a/Mathlib/RingTheory/Polynomial/Bernstein.lean b/Mathlib/RingTheory/Polynomial/Bernstein.lean index 108a392e18f23..5e3832422db3a 100644 --- a/Mathlib/RingTheory/Polynomial/Bernstein.lean +++ b/Mathlib/RingTheory/Polynomial/Bernstein.lean @@ -170,9 +170,7 @@ theorem iterate_derivative_at_0 (n ν : ℕ) : eval_natCast, Function.comp_apply, Function.iterate_succ, ascPochhammer_succ_left] obtain rfl | h'' := ν.eq_zero_or_pos · simp - · have : n - 1 - (ν - 1) = n - ν := by - rw [gt_iff_lt, ← Nat.succ_le_iff] at h'' - rw [← tsub_add_eq_tsub_tsub, add_comm, tsub_add_cancel_of_le h''] + · have : n - 1 - (ν - 1) = n - ν := by omega rw [this, ascPochhammer_eval_succ] rw_mod_cast [tsub_add_cancel_of_le (h'.trans n.pred_le)] · simp only [not_le] at h diff --git a/Mathlib/RingTheory/PowerSeries/Basic.lean b/Mathlib/RingTheory/PowerSeries/Basic.lean index 00d42f27bd55b..6baea6db2fd4a 100644 --- a/Mathlib/RingTheory/PowerSeries/Basic.lean +++ b/Mathlib/RingTheory/PowerSeries/Basic.lean @@ -590,7 +590,7 @@ lemma coeff_one_pow (n : ℕ) (φ : R⟦X⟧) : rcases Nat.eq_zero_or_pos n with (rfl | hn) · simp induction n with - | zero => by_contra; omega + | zero => omega | succ n' ih => have h₁ (m : ℕ) : φ ^ (m + 1) = φ ^ m * φ := by exact rfl have h₂ : Finset.antidiagonal 1 = {(0, 1), (1, 0)} := by exact rfl From 9adbe92f24aae8678652260c3bcc556999fd22a0 Mon Sep 17 00:00:00 2001 From: "Tristan F.-R." Date: Thu, 26 Dec 2024 07:24:15 +0000 Subject: [PATCH 003/189] docs: fix typos across repository (#20239) Supercedes https://github.com/leanprover-community/mathlib4/pull/20236. Hopefully this stops a few typos pull requests later. Uses `crates-ci/typos`, with the following config: ```toml [default.extend-words] # keywords & abbreviations pn = "pn" hae = "hae" hom = "hom" mor = "mor" thm = "thm" tpos = "tpos" hpe = "hpe" fle = "fle" eis = "eis" iy = "iy" hge = "hge" alls = "alls" struc = "struc" onl = "onl" hav = "hav" hax = "hax" ba = "ba" ein = "ein" adn = "adn" loopup = "loopup" hsa = "hsa" dum = "dum" mapp = "mapp" nd = "nd" ist = "ist" haa = "haa" ot = "ot" comon = "comon" ihs = "ihs" thr = "thr" projectives = "projectives" hda = "hda" ane = "ane" hsi = "hsi" minimals = "minimals" ninj = "ninj" ue = "ue" tne = "tne" lits = "lits" ine = "ine" tre = "tre" fo = "fo" nam = "nam" larg = "larg" reord = "reord" nin = "nin" toi = "toi" ands = "ands" hte = "hte" strat = "strat" hink = "hink" lastr = "lastr" wth = "wth" iin = "iin" # deprecations based on typos continous = "continous" # names and non-english words falso = "falso" yoh = "yoh" tak = "tak" lacker = "lacker" fonctions = "fonctions" calle = "calle" gool = "gool" sur = "sur" periodes = "periodes" waring = "waring" espaces = "espaces" yau = "yau" topologie = "topologie" claus = "claus" parth = "parth" ines = "ines" fike = "fike" tunnell = "tunnell" groupes = "groupes" htmp = "htmp" ket = "ket" clos = "clos" abd = "abd" # typo that i'm worried about breaking conciousness = "conciousness" # part of a test assum = "assum" # this might just be an abbrv but it's in a tactic and not used anywhere else [files] extend-exclude = ["docs/references.bib"] ``` I've dropped this config from this PR, but these typos ran across the repository - it may be beneficial for there to be CI to stop these unnecessary theorem deprecations. However, as the config also indicates, this captures a lot of false positives from the various abbreviations used. Co-authored-by: Tristan F. Co-authored-by: Tristan F.-R. --- Archive/Wiedijk100Theorems/BallotProblem.lean | 4 +-- ...DiscreteTopologyNonDiscreteUniformity.lean | 8 +++--- Mathlib/Algebra/Group/Defs.lean | 2 +- .../Module/Presentation/DirectSum.lean | 2 +- Mathlib/Algebra/Module/SnakeLemma.lean | 2 +- .../Algebra/Module/Submodule/Invariant.lean | 2 +- .../Algebra/MvPolynomial/SchwartzZippel.lean | 2 +- .../Algebra/Order/Group/Unbundled/Abs.lean | 4 +-- Mathlib/Algebra/Tropical/BigOperators.lean | 2 +- .../Cover/MorphismProperty.lean | 2 +- .../Morphisms/SurjectiveOnStalks.lean | 2 +- Mathlib/AlgebraicGeometry/Over.lean | 2 +- .../ProjectiveSpectrum/Basic.lean | 2 +- Mathlib/AlgebraicGeometry/Sites/Small.lean | 2 +- .../Analysis/BoxIntegral/UnitPartition.lean | 2 +- .../CStarAlgebra/ApproximateUnit.lean | 4 +-- Mathlib/Analysis/Normed/Order/Hom/Ultra.lean | 2 +- Mathlib/CategoryTheory/Comma/OverClass.lean | 2 +- Mathlib/CategoryTheory/Limits/Final.lean | 8 +++--- .../CategoryTheory/Limits/FinallySmall.lean | 2 +- .../Limits/FunctorCategory/Basic.lean | 2 +- .../Limits/Indization/IndObject.lean | 2 +- .../Limits/Preserves/Opposites.lean | 4 ++- .../Limits/Shapes/IsTerminal.lean | 2 +- .../Sites/EqualizerSheafCondition.lean | 4 +-- Mathlib/CategoryTheory/Yoneda.lean | 6 ++--- .../Condensed/Discrete/Characterization.lean | 16 +++++++++--- Mathlib/Data/Finsupp/SMulWithZero.lean | 2 +- Mathlib/Data/Matrix/Block.lean | 2 +- Mathlib/Data/Matrix/Invertible.lean | 4 +-- Mathlib/Data/Matrix/Rank.lean | 2 +- Mathlib/Data/PFunctor/Univariate/M.lean | 4 ++- Mathlib/Data/Real/EReal.lean | 2 +- Mathlib/FieldTheory/IsAlgClosed/Basic.lean | 2 +- .../Geometry/Manifold/Instances/Sphere.lean | 2 +- Mathlib/GroupTheory/Nilpotent.lean | 9 ++++--- Mathlib/GroupTheory/Sylow.lean | 2 +- Mathlib/GroupTheory/Transfer.lean | 2 +- .../LinearAlgebra/AffineSpace/AffineMap.lean | 4 +-- Mathlib/LinearAlgebra/Eigenspace/Pi.lean | 2 +- .../LinearAlgebra/Matrix/SchurComplement.lean | 2 +- Mathlib/LinearAlgebra/PerfectPairing.lean | 2 +- Mathlib/MeasureTheory/Covering/Vitali.lean | 25 +++++++++++++------ .../MeasureTheory/Integral/SetIntegral.lean | 2 +- .../MeasurableSpace/CountablyGenerated.lean | 2 +- Mathlib/NumberTheory/Cyclotomic/Three.lean | 2 +- Mathlib/NumberTheory/LSeries/PrimesInAP.lean | 6 ++--- .../NumberTheory/NumberField/Completion.lean | 2 +- Mathlib/NumberTheory/Padics/MahlerBasis.lean | 2 +- .../RamificationInertia/Basic.lean | 2 +- Mathlib/Order/Grade.lean | 4 +-- Mathlib/Order/KrullDimension.lean | 2 +- Mathlib/Order/SuccPred/WithBot.lean | 2 +- Mathlib/Probability/Independence/Kernel.lean | 2 +- .../HomogeneousLocalization.lean | 2 +- Mathlib/RingTheory/Lasker.lean | 4 +-- Mathlib/RingTheory/LaurentSeries.lean | 4 +-- .../MvPolynomial/EulerIdentity.lean | 2 +- Mathlib/RingTheory/MvPowerSeries/Order.lean | 2 +- .../RingTheory/PowerSeries/PiTopology.lean | 2 +- .../RootsOfUnity/AlgebraicallyClosed.lean | 2 +- .../Valuation/ValuationSubring.lean | 2 +- Mathlib/Tactic/Algebraize.lean | 4 +-- Mathlib/Tactic/FunProp/Core.lean | 2 +- Mathlib/Tactic/FunProp/Theorems.lean | 2 +- Mathlib/Tactic/ReduceModChar.lean | 2 +- Mathlib/Tactic/UnsetOption.lean | 2 +- Mathlib/Tactic/Variable.lean | 2 +- Mathlib/Topology/Connected/Clopen.lean | 2 +- Mathlib/Topology/Order/OrderClosed.lean | 4 +-- Mathlib/Topology/UniformSpace/Dini.lean | 2 +- MathlibTest/algebraize.lean | 8 +++--- MathlibTest/zmod.lean | 2 +- scripts/add_deprecations.sh | 2 +- scripts/late_importers.sh | 2 +- 75 files changed, 136 insertions(+), 110 deletions(-) diff --git a/Archive/Wiedijk100Theorems/BallotProblem.lean b/Archive/Wiedijk100Theorems/BallotProblem.lean index 1476d785b39a2..9ee2b014457bb 100644 --- a/Archive/Wiedijk100Theorems/BallotProblem.lean +++ b/Archive/Wiedijk100Theorems/BallotProblem.lean @@ -165,9 +165,9 @@ theorem disjoint_bits (p q : ℕ) : open MeasureTheory.Measure -private def measureableSpace_list_int : MeasurableSpace (List ℤ) := ⊤ +private def measurableSpace_list_int : MeasurableSpace (List ℤ) := ⊤ -attribute [local instance] measureableSpace_list_int +attribute [local instance] measurableSpace_list_int private theorem measurableSingletonClass_list_int : MeasurableSingletonClass (List ℤ) := { measurableSet_singleton := fun _ => trivial } diff --git a/Counterexamples/DiscreteTopologyNonDiscreteUniformity.lean b/Counterexamples/DiscreteTopologyNonDiscreteUniformity.lean index 572f6fc61a1b3..1da3cf068ae49 100644 --- a/Counterexamples/DiscreteTopologyNonDiscreteUniformity.lean +++ b/Counterexamples/DiscreteTopologyNonDiscreteUniformity.lean @@ -20,7 +20,7 @@ element of the lattice of uniformities on a type (see `bot_uniformity`). The theorem `discreteTopology_of_discrete_uniformity` shows that the topology induced by the discrete uniformity is the discrete one, but it is well-known that the converse might not hold in -general, along the lines of the above discussion. We explicitely produce a metric and a uniform +general, along the lines of the above discussion. We explicitly produce a metric and a uniform structure on a space (on `ℕ`, actually) that are not discrete, yet induce the discrete topology. To check that a certain uniformity is not discrete, recall that once a type `α` is endowed with a @@ -31,7 +31,7 @@ declaration `UniformSpace.DiscreteUnif.eq_const_of_cauchy` in Mathlib. A special case of this result is the intuitive observation that a sequence `a : ℕ → ℕ` can be a Cauchy sequence if and only if it is eventually constant: when claiming this equivalence, one is -implicitely endowing `ℕ` with the metric inherited from `ℝ`, that induces the discrete uniformity +implicitly endowing `ℕ` with the metric inherited from `ℝ`, that induces the discrete uniformity on `ℕ`. Hence, the intuition suggesting that a Cauchy sequence, whose terms are "closer and closer to each other", valued in `ℕ` must be eventually constant for *topological* reasons, namely the fact that `ℕ` is a discrete topological space, is *wrong* in the @@ -41,7 +41,7 @@ the identity `id : ℕ → ℕ` is Cauchy, then the uniformity is certainly *not ## The counterexamples We produce two counterexamples: in the first section `Metric` we construct a metric and in the -second section `SetPointUniformity` we construct a uniformity, explicitely as a filter on `ℕ × ℕ`. +second section `SetPointUniformity` we construct a uniformity, explicitly as a filter on `ℕ × ℕ`. They basically coincide, and the difference of the examples lies in their flavours. ### The metric @@ -232,7 +232,7 @@ lemma mem_counterBasis_iff (S : Set (ℕ × ℕ)) : /-- The "crude" uniform structure, without topology, simply as a the filter generated by `Basis` and satisfying the axioms for being a uniformity. We later extract the topology `counterTopology` -generated by it and bundle `counterCoreUniformity` and `counterTopology` in a uniform strucutre +generated by it and bundle `counterCoreUniformity` and `counterTopology` in a uniform structure on `ℕ`, proving in passing that `counterTopology = ⊥`-/ def counterCoreUniformity : UniformSpace.Core ℕ := by apply UniformSpace.Core.mkOfBasis counterBasis <;> diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index 13e7ae50e5308..ab65c89102d88 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -409,7 +409,7 @@ the `npow` field when defining multiplicative objects. /-- Exponentiation by repeated squaring. -/ @[to_additive "Scalar multiplication by repeated self-addition, -the additive version of exponentation by repeated squaring."] +the additive version of exponentiation by repeated squaring."] def npowBinRec {M : Type*} [One M] [Mul M] (k : ℕ) : M → M := npowBinRec.go k 1 where diff --git a/Mathlib/Algebra/Module/Presentation/DirectSum.lean b/Mathlib/Algebra/Module/Presentation/DirectSum.lean index 387bdbda99e01..286e36d10e7b6 100644 --- a/Mathlib/Algebra/Module/Presentation/DirectSum.lean +++ b/Mathlib/Algebra/Module/Presentation/DirectSum.lean @@ -13,7 +13,7 @@ import Mathlib.Data.Finsupp.ToDFinsupp If `M : ι → Type _` is a family of `A`-modules, then the data of a presentation of each `M i`, we obtain a presentation of the module `⨁ i, M i`. In particular, from a presentation of an `A`-module `M`, we get -a presention of `ι →₀ M`. +a presentation of `ι →₀ M`. -/ diff --git a/Mathlib/Algebra/Module/SnakeLemma.lean b/Mathlib/Algebra/Module/SnakeLemma.lean index 4ddf484287f4e..3c3aea5de4924 100644 --- a/Mathlib/Algebra/Module/SnakeLemma.lean +++ b/Mathlib/Algebra/Module/SnakeLemma.lean @@ -13,7 +13,7 @@ The snake lemma is proven in `Algebra/Homology/ShortComplex/SnakeLemma.lean` for categories, but for definitional equality and universe issues we reprove them here for modules. ## Main results -- `SnakeLemma.δ`: The connecting homomorphism guranteed by the snake lemma. +- `SnakeLemma.δ`: The connecting homomorphism guaranteed by the snake lemma. - `SnakeLemma.exact_δ_left`: The connecting homomorphism is exact on the right. - `SnakeLemma.exact_δ_right`: The connecting homomorphism is exact on the left. diff --git a/Mathlib/Algebra/Module/Submodule/Invariant.lean b/Mathlib/Algebra/Module/Submodule/Invariant.lean index cfb3ce8b980af..42f48782a8d57 100644 --- a/Mathlib/Algebra/Module/Submodule/Invariant.lean +++ b/Mathlib/Algebra/Module/Submodule/Invariant.lean @@ -10,7 +10,7 @@ import Mathlib.Order.Sublattice # The lattice of invariant submodules In this file we defined the type `Module.End.invtSubmodule`, associated to a linear endomorphism of -a module. Its utilty stems primarily from those occasions on which we wish to take advantage of the +a module. Its utility stems primarily from those occasions on which we wish to take advantage of the lattice structure of invariant submodules. See also `Mathlib.Algebra.Polynomial.Module.AEval`. diff --git a/Mathlib/Algebra/MvPolynomial/SchwartzZippel.lean b/Mathlib/Algebra/MvPolynomial/SchwartzZippel.lean index 1d5732412ed20..9b5e8cdc2e08b 100644 --- a/Mathlib/Algebra/MvPolynomial/SchwartzZippel.lean +++ b/Mathlib/Algebra/MvPolynomial/SchwartzZippel.lean @@ -57,7 +57,7 @@ namespace MvPolynomial variable {R : Type*} [CommRing R] [IsDomain R] [DecidableEq R] -- A user should be able to provide `hp` as a named argument --- regardless of whether one has used pattern-maching or induction to prove the lemma. +-- regardless of whether one has used pattern-matching or induction to prove the lemma. set_option linter.unusedVariables false in /-- The **Schwartz-Zippel lemma** diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean b/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean index bb4c1570f4cbc..56042ed6f46bf 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean @@ -38,7 +38,7 @@ macro:max atomic("|" noWs) a:term noWs "|ₘ" : term => `(mabs $a) macro:max atomic("|" noWs) a:term noWs "|" : term => `(abs $a) /-- Unexpander for the notation `|a|ₘ` for `mabs a`. -Tries to add discretionary parentheses in unparseable cases. -/ +Tries to add discretionary parentheses in unparsable cases. -/ @[app_unexpander abs] def mabs.unexpander : Lean.PrettyPrinter.Unexpander | `($_ $a) => @@ -48,7 +48,7 @@ def mabs.unexpander : Lean.PrettyPrinter.Unexpander | _ => throw () /-- Unexpander for the notation `|a|` for `abs a`. -Tries to add discretionary parentheses in unparseable cases. -/ +Tries to add discretionary parentheses in unparsable cases. -/ @[app_unexpander abs] def abs.unexpander : Lean.PrettyPrinter.Unexpander | `($_ $a) => diff --git a/Mathlib/Algebra/Tropical/BigOperators.lean b/Mathlib/Algebra/Tropical/BigOperators.lean index e0c31e19c6c04..f6c25370a25b6 100644 --- a/Mathlib/Algebra/Tropical/BigOperators.lean +++ b/Mathlib/Algebra/Tropical/BigOperators.lean @@ -23,7 +23,7 @@ collection of linear functions. ## Implementation notes -No concrete (semi)ring is used here, only ones with inferrable order/lattice structure, to support +No concrete (semi)ring is used here, only ones with inferable order/lattice structure, to support `Real`, `Rat`, `EReal`, and others (`ERat` is not yet defined). Minima over `List α` are defined as producing a value in `WithTop α` so proofs about lists do not diff --git a/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean b/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean index 53a9c4dfd2092..87849e46136ca 100644 --- a/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean +++ b/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean @@ -89,7 +89,7 @@ def Cover.mkOfCovers (J : Type*) (obj : J → Scheme.{u}) (map : (j : J) → obj covers x := (covers x).choose_spec map_prop := map_prop -/-- Turn a `P`-cover into a `Q`-cover by showing that the components satisfiy `Q`. -/ +/-- Turn a `P`-cover into a `Q`-cover by showing that the components satisfy `Q`. -/ def Cover.changeProp (Q : MorphismProperty Scheme.{u}) (𝒰 : X.Cover P) (h : ∀ j, Q (𝒰.map j)) : X.Cover Q where J := 𝒰.J diff --git a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean index a39842eb7658c..b71393e088b97 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean @@ -10,7 +10,7 @@ import Mathlib.Topology.LocalAtTarget /-! # Morphisms surjective on stalks -We define the classe of morphisms between schemes that are surjective on stalks. +We define the class of morphisms between schemes that are surjective on stalks. We show that this class is stable under composition and base change. We also show that (`AlgebraicGeometry.SurjectiveOnStalks.isEmbedding_pullback`) diff --git a/Mathlib/AlgebraicGeometry/Over.lean b/Mathlib/AlgebraicGeometry/Over.lean index 15349f1b13a3e..6d0aabc1013fd 100644 --- a/Mathlib/AlgebraicGeometry/Over.lean +++ b/Mathlib/AlgebraicGeometry/Over.lean @@ -33,7 +33,7 @@ protected abbrev Over (X S : Scheme.{u}) := OverClass X S /-- `X.CanonicallyOver S` is the typeclass containing the data of a structure morphism `X ↘ S : X ⟶ S`, -and that `S` is (uniquely) inferrable from the structure of `X`. +and that `S` is (uniquely) inferable from the structure of `X`. -/ abbrev CanonicallyOver := CanonicallyOverClass X S diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean index 645435b9424b1..16327672ba49a 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean @@ -250,7 +250,7 @@ lemma pullbackAwayιIso_inv_snd : rw [← pullbackAwayιIso_hom_SpecMap_awayMap_right, Iso.inv_hom_id_assoc] open TopologicalSpace.Opens in -/-- Given a family of homogeneous elements `f` of positive degree that spans the irrelavent ideal, +/-- Given a family of homogeneous elements `f` of positive degree that spans the irrelevant ideal, `Spec (A_f)₀ ⟶ Proj A` forms an affine open cover of `Proj A`. -/ noncomputable def openCoverOfISupEqTop {ι : Type*} (f : ι → A) {m : ι → ℕ} diff --git a/Mathlib/AlgebraicGeometry/Sites/Small.lean b/Mathlib/AlgebraicGeometry/Sites/Small.lean index 5e5bfcaf299da..71d02013f5603 100644 --- a/Mathlib/AlgebraicGeometry/Sites/Small.lean +++ b/Mathlib/AlgebraicGeometry/Sites/Small.lean @@ -162,7 +162,7 @@ abbrev smallGrothendieckTopology : GrothendieckTopology (P.Over ⊤ S) := variable [Q.IsStableUnderBaseChange] [Q.HasOfPostcompProperty Q] -/-- The pretopology defind on the subcategory of `S`-schemes satisfying `Q` where coverings +/-- The pretopology defined on the subcategory of `S`-schemes satisfying `Q` where coverings are given by `P`-coverings in `S`-schemes satisfying `Q`. The most common case is `P = Q`. In this case, this is simply surjective families in `S`-schemes with `P`. -/ diff --git a/Mathlib/Analysis/BoxIntegral/UnitPartition.lean b/Mathlib/Analysis/BoxIntegral/UnitPartition.lean index 17d078e1e0df9..35e8c92b4a7a5 100644 --- a/Mathlib/Analysis/BoxIntegral/UnitPartition.lean +++ b/Mathlib/Analysis/BoxIntegral/UnitPartition.lean @@ -17,7 +17,7 @@ vectors `ν : ι → ℤ`. Let `B` be a `BoxIntegral`. A `unitPartition.box` is admissible for `B` (more precisely its index is admissible) if it is contained in `B`. There are finitely many admissible `unitPartition.box` for -`B` and thus we can form the corresponing tagged prepartition, see +`B` and thus we can form the corresponding tagged prepartition, see `BoxIntegral.unitPartition.prepartition` (note that each `unitPartition.box` comes with its tag situated at its "upper most" vertex). If `B` satisfies `hasIntegralVertices`, that is its vertices are in `ι → ℤ`, then the corresponding prepartition is actually a partition. diff --git a/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean b/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean index 49852fd88a32f..2f1dddf941c77 100644 --- a/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean +++ b/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean @@ -188,8 +188,8 @@ lemma hasBasis_approximateUnit : (approximateUnit A).HasBasis (fun x : A ↦ 0 ≤ x ∧ ‖x‖ < 1) ({x | · ≤ x} ∩ closedBall 0 1) := isBasis_nonneg_sections A |>.hasBasis.inf_principal (closedBall 0 1) -/-- This is a common reasoning sequence in C⋆-algebra theory. If `0 ≤ x ≤ y ≤ 1`, then the norm -of `z - y * z` is controled by the norm of `star z * (1 - x) * z`, which is advantageous because the +/-- This is a common reasoning sequence in C⋆-algebra theory. If `0 ≤ x ≤ y ≤ 1`, then the norm of +`z - y * z` is controlled by the norm of `star z * (1 - x) * z`, which is advantageous because the latter is nonnegative. This is a key step in establishing the existence of an increasing approximate unit in general C⋆-algebras. -/ lemma nnnorm_sub_mul_self_le {A : Type*} [CStarAlgebra A] [PartialOrder A] [StarOrderedRing A] diff --git a/Mathlib/Analysis/Normed/Order/Hom/Ultra.lean b/Mathlib/Analysis/Normed/Order/Hom/Ultra.lean index b28160bad854a..37f55025acaf7 100644 --- a/Mathlib/Analysis/Normed/Order/Hom/Ultra.lean +++ b/Mathlib/Analysis/Normed/Order/Hom/Ultra.lean @@ -23,7 +23,7 @@ the argument is an autoparam that resolves by definitional equality when using t variable {F α : Type*} [FunLike F α ℝ] /-- Proves that when a `SeminormedAddGroup` structure is constructed from an -`AddGroupSeminormClass` that satifies `IsNonarchimedean`, the group has an `IsUltrametricDist`. -/ +`AddGroupSeminormClass` that satisfies `IsNonarchimedean`, the group has an `IsUltrametricDist`. -/ lemma AddGroupSeminormClass.isUltrametricDist [AddGroup α] [AddGroupSeminormClass F α ℝ] [inst : Dist α] {f : F} (hna : IsNonarchimedean f) (hd : inst = (AddGroupSeminormClass.toSeminormedAddGroup f).toDist := by rfl) : diff --git a/Mathlib/CategoryTheory/Comma/OverClass.lean b/Mathlib/CategoryTheory/Comma/OverClass.lean index a6b1fc35caeb2..0b1331517387d 100644 --- a/Mathlib/CategoryTheory/Comma/OverClass.lean +++ b/Mathlib/CategoryTheory/Comma/OverClass.lean @@ -58,7 +58,7 @@ initialize_simps_projections OverClass (hom → over) /-- `X.CanonicallyOverClass S` is the typeclass containing the data of a structure morphism `X ↘ S : X ⟶ S`, -and that `S` is (uniquely) inferrable from the structure of `X`. +and that `S` is (uniquely) inferable from the structure of `X`. -/ class CanonicallyOverClass (X : C) (S : semiOutParam C) extends OverClass X S where diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index 6719a13bc510a..f704225f03972 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -796,7 +796,7 @@ theorem initial_comp_equivalence [Initial F] [IsEquivalence G] : Initial (F ⋙ theorem final_equivalence_comp [IsEquivalence F] [Final G] : Final (F ⋙ G) where out d := isConnected_of_equivalent (StructuredArrow.pre d F G).asEquivalence.symm -/-- See also the strictly more general `inital_comp` below. -/ +/-- See also the strictly more general `initial_comp` below. -/ theorem initial_equivalence_comp [IsEquivalence F] [Initial G] : Initial (F ⋙ G) where out d := isConnected_of_equivalent (CostructuredArrow.pre F G d).asEquivalence.symm @@ -941,7 +941,8 @@ theorem IsFiltered.of_final (F : C ⥤ D) [Final F] [IsFiltered C] : IsFiltered /-- Initial functors preserve cofilteredness. This can be seen as a generalization of `IsCofiltered.of_left_adjoint` (which states that left -adjoints preserve cofilteredness), as right adjoints are always initial, see `intial_of_adjunction`. +adjoints preserve cofilteredness), as right adjoints are always initial, +see `initial_of_adjunction`. -/ theorem IsCofilteredOrEmpty.of_initial (F : C ⥤ D) [Initial F] [IsCofilteredOrEmpty C] : IsCofilteredOrEmpty D := @@ -951,7 +952,8 @@ theorem IsCofilteredOrEmpty.of_initial (F : C ⥤ D) [Initial F] [IsCofilteredOr /-- Initial functors preserve cofilteredness. This can be seen as a generalization of `IsCofiltered.of_left_adjoint` (which states that left -adjoints preserve cofilteredness), as right adjoints are always initial, see `intial_of_adjunction`. +adjoints preserve cofilteredness), as right adjoints are always initial, +see `initial_of_adjunction`. -/ theorem IsCofiltered.of_initial (F : C ⥤ D) [Initial F] [IsCofiltered C] : IsCofiltered D := have : IsFiltered Dᵒᵖ := IsFiltered.of_final F.op diff --git a/Mathlib/CategoryTheory/Limits/FinallySmall.lean b/Mathlib/CategoryTheory/Limits/FinallySmall.lean index 3613b80d171c2..7e7c96252359a 100644 --- a/Mathlib/CategoryTheory/Limits/FinallySmall.lean +++ b/Mathlib/CategoryTheory/Limits/FinallySmall.lean @@ -160,7 +160,7 @@ section WeaklyInitial variable (J : Type u) [Category.{v} J] -/-- The converse is true if `J` is cofiltered, see `intiallySmall_of_small_weakly_initial_set`. -/ +/-- The converse is true if `J` is cofiltered, see `initiallySmall_of_small_weakly_initial_set`. -/ theorem InitiallySmall.exists_small_weakly_initial_set [InitiallySmall.{w} J] : ∃ (s : Set J) (_ : Small.{w} s), ∀ i, ∃ j ∈ s, Nonempty (j ⟶ i) := by refine ⟨Set.range (fromInitialModel J).obj, inferInstance, fun i => ?_⟩ diff --git a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean index b9d09a13f83f1..c22fca4cbe509 100644 --- a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean @@ -197,7 +197,7 @@ instance functorCategoryHasColimitsOfSize [HasColimitsOfSize.{v₁, u₁} C] : HasColimitsOfSize.{v₁, u₁} (K ⥤ C) where has_colimits_of_shape := fun _ _ => inferInstance -instance hasLimitCompEvalution (F : J ⥤ K ⥤ C) (k : K) [HasLimit (F.flip.obj k)] : +instance hasLimitCompEvaluation (F : J ⥤ K ⥤ C) (k : K) [HasLimit (F.flip.obj k)] : HasLimit (F ⋙ (evaluation _ _).obj k) := hasLimitOfIso (F := F.flip.obj k) (Iso.refl _) diff --git a/Mathlib/CategoryTheory/Limits/Indization/IndObject.lean b/Mathlib/CategoryTheory/Limits/Indization/IndObject.lean index d428ddcc7ae8b..4594f661240c5 100644 --- a/Mathlib/CategoryTheory/Limits/Indization/IndObject.lean +++ b/Mathlib/CategoryTheory/Limits/Indization/IndObject.lean @@ -57,7 +57,7 @@ structure IndObjectPresentation (A : Cᵒᵖ ⥤ Type v) where F : I ⥤ C /-- Use `IndObjectPresentation.cocone` instead. -/ ι : F ⋙ yoneda ⟶ (Functor.const I).obj A - /-- Use `IndObjectPresenation.coconeIsColimit` instead. -/ + /-- Use `IndObjectPresentation.coconeIsColimit` instead. -/ isColimit : IsColimit (Cocone.mk A ι) namespace IndObjectPresentation diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean b/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean index 3c6706cf053d9..5a089c8cbb3e7 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean @@ -319,9 +319,11 @@ lemma preservesLimits_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimits F] : Pre preservesLimitsOfShape {_} _ := preservesLimitsOfShape_unop _ _ /-- If `F : C ⥤ D` preserves limits, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits. -/ -lemma perservesColimits_op (F : C ⥤ D) [PreservesLimits F] : PreservesColimits F.op where +lemma preservesColimits_op (F : C ⥤ D) [PreservesLimits F] : PreservesColimits F.op where preservesColimitsOfShape {_} _ := preservesColimitsOfShape_op _ _ +@[deprecated (since := "2024-12-25")] alias perservesColimits_op := preservesColimits_op + /-- If `F : C ⥤ Dᵒᵖ` preserves limits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits. -/ lemma preservesColimits_leftOp (F : C ⥤ Dᵒᵖ) [PreservesLimits F] : PreservesColimits F.leftOp where preservesColimitsOfShape {_} _ := preservesColimitsOfShape_leftOp _ _ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean b/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean index 7ceb4a6aeb551..de6ee324b9d3d 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean @@ -14,7 +14,7 @@ import Mathlib.CategoryTheory.Category.Preorder In this file we define the predicates `IsTerminal` and `IsInitial` as well as the class `InitialMonoClass`. -The classes `HasTerminal` and `HasInitial` and the associated notations for terminal and inital +The classes `HasTerminal` and `HasInitial` and the associated notations for terminal and initial objects are defined in `Terminal.lean`. ## References diff --git a/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean b/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean index 1df38a2c7b41e..2679d293f850b 100644 --- a/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean +++ b/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean @@ -256,7 +256,7 @@ variable {B : C} {I : Type} (X : I → C) (π : (i : I) → X i ⟶ B) /-- The middle object of the fork diagram of . -The difference between this and `Equalizer.FirstObj P (ofArrows X π)` arrises if the family of +The difference between this and `Equalizer.FirstObj P (ofArrows X π)` arises if the family of arrows `π` contains duplicates. The `Presieve.ofArrows` doesn't see those. -/ def FirstObj : Type w := ∏ᶜ (fun i ↦ P.obj (op (X i))) @@ -270,7 +270,7 @@ lemma FirstObj.ext (z₁ z₂ : FirstObj P X) (h : ∀ i, (Pi.π _ i : FirstObj /-- The rightmost object of the fork diagram of https://stacks.math.columbia.edu/tag/00VM. -The difference between this and `Equalizer.Presieve.SecondObj P (ofArrows X π)` arrises if the +The difference between this and `Equalizer.Presieve.SecondObj P (ofArrows X π)` arises if the family of arrows `π` contains duplicates. The `Presieve.ofArrows` doesn't see those. -/ def SecondObj : Type w := diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index 08ebb0e9e9c8f..05d0c28efabb3 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -252,7 +252,7 @@ def CorepresentableBy.toIso {F : C ⥤ Type v₁} {X : C} (e : F.Corepresentable coyoneda.obj (op X) ≅ F := corepresentableByEquiv e -/-- A functor `F : Cᵒᵖ ⥤ Type v` is representable if there is oan bject `Y` with a structure +/-- A functor `F : Cᵒᵖ ⥤ Type v` is representable if there is an object `Y` with a structure `F.RepresentableBy Y`, i.e. there is a natural bijection `(X ⟶ Y) ≃ F.obj (op X)`, which may also be rephrased as a natural isomorphism `yoneda.obj X ≅ F` when `Category.{v} C`. @@ -267,7 +267,7 @@ lemma RepresentableBy.isRepresentable {F : Cᵒᵖ ⥤ Type v} {Y : C} (e : F.Re F.IsRepresentable where has_representation := ⟨Y, ⟨e⟩⟩ -/-- Alternative constructure for `F.IsRepresentable`, which takes as an input an +/-- Alternative constructor for `F.IsRepresentable`, which takes as an input an isomorphism `yoneda.obj X ≅ F`. -/ lemma IsRepresentable.mk' {F : Cᵒᵖ ⥤ Type v₁} {X : C} (e : yoneda.obj X ≅ F) : F.IsRepresentable := @@ -289,7 +289,7 @@ lemma CorepresentableBy.isCorepresentable {F : C ⥤ Type v} {X : C} (e : F.Core F.IsCorepresentable where has_corepresentation := ⟨X, ⟨e⟩⟩ -/-- Alternative constructure for `F.IsCorepresentable`, which takes as an input an +/-- Alternative constructor for `F.IsCorepresentable`, which takes as an input an isomorphism `coyoneda.obj (op X) ≅ F`. -/ lemma IsCorepresentable.mk' {F : C ⥤ Type v₁} {X : C} (e : coyoneda.obj (op X) ≅ F) : F.IsCorepresentable := diff --git a/Mathlib/Condensed/Discrete/Characterization.lean b/Mathlib/Condensed/Discrete/Characterization.lean index 842f3c46d0442..2a4d5db7d2127 100644 --- a/Mathlib/Condensed/Discrete/Characterization.lean +++ b/Mathlib/Condensed/Discrete/Characterization.lean @@ -50,7 +50,7 @@ namespace CondensedSet open CompHausLike.LocallyConstant -lemma mem_locallyContant_essImage_of_isColimit_mapCocone (X : CondensedSet.{u}) +lemma mem_locallyConstant_essImage_of_isColimit_mapCocone (X : CondensedSet.{u}) (h : ∀ S : Profinite.{u}, IsColimit <| (profiniteToCompHaus.op ⋙ X.val).mapCocone S.asLimitCone.op) : X ∈ CondensedSet.LocallyConstant.functor.essImage := by @@ -60,6 +60,10 @@ lemma mem_locallyContant_essImage_of_isColimit_mapCocone (X : CondensedSet.{u}) Condensed.isoLocallyConstantOfIsColimit _ h exact ⟨_, ⟨e.functor.preimageIso ((sheafToPresheaf _ _).preimageIso i.symm)⟩⟩ +@[deprecated (since := "2024-12-25")] +alias mem_locallyContant_essImage_of_isColimit_mapCocone := + mem_locallyConstant_essImage_of_isColimit_mapCocone + /-- `CondensedSet.LocallyConstant.functor` is left adjoint to the forgetful functor from condensed sets to sets. @@ -96,7 +100,7 @@ theorem isDiscrete_tfae (X : CondensedSet.{u}) : (coherentTopology CompHaus) profiniteToCompHaus Profinite.isTerminalPUnit CompHaus.isTerminalPUnit _).symm tfae_have 7 → 4 := fun h ↦ - mem_locallyContant_essImage_of_isColimit_mapCocone X (fun S ↦ (h S).some) + mem_locallyConstant_essImage_of_isColimit_mapCocone X (fun S ↦ (h S).some) tfae_have 4 → 7 := fun ⟨Y, ⟨i⟩⟩ S ↦ ⟨IsColimit.mapCoconeEquiv (isoWhiskerLeft profiniteToCompHaus.op ((sheafToPresheaf _ _).mapIso i)) @@ -173,7 +177,7 @@ end LightCondensed namespace LightCondSet -lemma mem_locallyContant_essImage_of_isColimit_mapCocone (X : LightCondSet.{u}) +lemma mem_locallyConstant_essImage_of_isColimit_mapCocone (X : LightCondSet.{u}) (h : ∀ S : LightProfinite.{u}, IsColimit <| X.val.mapCocone (coconeRightOpOfCone S.asLimitCone)) : X ∈ LightCondSet.LocallyConstant.functor.essImage := by @@ -181,6 +185,10 @@ lemma mem_locallyContant_essImage_of_isColimit_mapCocone (X : LightCondSet.{u}) LightCondensed.isoLocallyConstantOfIsColimit _ h exact ⟨_, ⟨((sheafToPresheaf _ _).preimageIso i.symm)⟩⟩ +@[deprecated (since := "2024-12-25")] +alias mem_locallyContant_essImage_of_isColimit_mapCocone := + mem_locallyConstant_essImage_of_isColimit_mapCocone + /-- `LightCondSet.LocallyConstant.functor` is left adjoint to the forgetful functor from light condensed sets to sets. @@ -209,7 +217,7 @@ theorem isDiscrete_tfae (X : LightCondSet.{u}) : -- These `have` statements above shouldn't be needed, but they are. Sheaf.isConstant_iff_isIso_counit_app' _ LightProfinite.isTerminalPUnit adjunction X tfae_have 6 → 4 := fun h ↦ - mem_locallyContant_essImage_of_isColimit_mapCocone X (fun S ↦ (h S).some) + mem_locallyConstant_essImage_of_isColimit_mapCocone X (fun S ↦ (h S).some) tfae_have 4 → 6 := fun ⟨Y, ⟨i⟩⟩ S ↦ ⟨IsColimit.mapCoconeEquiv ((sheafToPresheaf _ _).mapIso i) (LightCondensed.isColimitLocallyConstantPresheafDiagram Y S)⟩ diff --git a/Mathlib/Data/Finsupp/SMulWithZero.lean b/Mathlib/Data/Finsupp/SMulWithZero.lean index e7e1adc5d920c..79dda0d1067c9 100644 --- a/Mathlib/Data/Finsupp/SMulWithZero.lean +++ b/Mathlib/Data/Finsupp/SMulWithZero.lean @@ -21,7 +21,7 @@ This file defines the pointwise scalar multiplication on `Finsupp`, assuming it This file is intermediate between `Finsupp.Defs` and `Finsupp.Module` in that it covers scalar multiplication but does not rely on the definition of `Module`. Scalar multiplication is needed to supply the `nsmul` (and `zsmul`) fields of (semi)ring structures which are fundamental for e.g. -`Polynomial`, so we want to keep the imports requied for the `Finsupp.smulZeroClass` instance +`Polynomial`, so we want to keep the imports required for the `Finsupp.smulZeroClass` instance reasonably light. This file is a `noncomputable theory` and uses classical logic throughout. diff --git a/Mathlib/Data/Matrix/Block.lean b/Mathlib/Data/Matrix/Block.lean index 15185b1ef2dbb..d19568fcf11b6 100644 --- a/Mathlib/Data/Matrix/Block.lean +++ b/Mathlib/Data/Matrix/Block.lean @@ -310,7 +310,7 @@ section Zero variable [Zero α] [Zero β] -/-- `Matrix.blockDiagonal M` turns a homogenously-indexed collection of matrices +/-- `Matrix.blockDiagonal M` turns a homogeneously-indexed collection of matrices `M : o → Matrix m n α'` into an `m × o`-by-`n × o` block matrix which has the entries of `M` along the diagonal and zero elsewhere. diff --git a/Mathlib/Data/Matrix/Invertible.lean b/Mathlib/Data/Matrix/Invertible.lean index 9a2edad88d15b..cb93d30186d50 100644 --- a/Mathlib/Data/Matrix/Invertible.lean +++ b/Mathlib/Data/Matrix/Invertible.lean @@ -17,8 +17,8 @@ in `LinearAlgebra/Matrix/NonsingularInverse.lean`. * `Matrix.invertibleConjTranspose` * `Matrix.invertibleTranspose` -* `Matrix.isUnit_conjTranpose` -* `Matrix.isUnit_tranpose` +* `Matrix.isUnit_conjTranspose` +* `Matrix.isUnit_transpose` -/ diff --git a/Mathlib/Data/Matrix/Rank.lean b/Mathlib/Data/Matrix/Rank.lean index 69228715dac42..e33a47769f203 100644 --- a/Mathlib/Data/Matrix/Rank.lean +++ b/Mathlib/Data/Matrix/Rank.lean @@ -165,7 +165,7 @@ section Field variable [Field R] -/-- The rank of a diagnonal matrix is the count of non-zero elements on its main diagonal -/ +/-- The rank of a diagonal matrix is the count of non-zero elements on its main diagonal -/ theorem rank_diagonal [Fintype m] [DecidableEq m] [DecidableEq R] (w : m → R) : (diagonal w).rank = Fintype.card {i // (w i) ≠ 0} := by rw [Matrix.rank, ← Matrix.toLin'_apply', Module.finrank, ← LinearMap.rank, diff --git a/Mathlib/Data/PFunctor/Univariate/M.lean b/Mathlib/Data/PFunctor/Univariate/M.lean index f82ab4ab17d8a..70b6e183f4a7e 100644 --- a/Mathlib/Data/PFunctor/Univariate/M.lean +++ b/Mathlib/Data/PFunctor/Univariate/M.lean @@ -73,7 +73,9 @@ def AllAgree (x : ∀ n, CofixA F n) := ∀ n, Agree (x n) (x (succ n)) @[simp] -theorem agree_trival {x : CofixA F 0} {y : CofixA F 1} : Agree x y := by constructor +theorem agree_trivial {x : CofixA F 0} {y : CofixA F 1} : Agree x y := by constructor + +@[deprecated (since := "2024-12-25")] alias agree_trival := agree_trivial theorem agree_children {n : ℕ} (x : CofixA F (succ n)) (y : CofixA F (succ n + 1)) {i j} (h₀ : HEq i j) (h₁ : Agree x y) : Agree (children' x i) (children' y j) := by diff --git a/Mathlib/Data/Real/EReal.lean b/Mathlib/Data/Real/EReal.lean index b2f392cdd17f4..66445295364a8 100644 --- a/Mathlib/Data/Real/EReal.lean +++ b/Mathlib/Data/Real/EReal.lean @@ -31,7 +31,7 @@ choice that `0 * x = x * 0 = 0` for any `x` (while the other cases are defined n This does not distribute with addition, as `⊥ = ⊥ + ⊤ = 1*⊥ + (-1)*⊥ ≠ (1 - 1) * ⊥ = 0 * ⊥ = 0`. Distributivity `x * (y + z) = x * y + x * z` is recovered in the case where either `0 ≤ x < ⊤`, see `Ereal.left_distrib_of_nonneg_of_ne_top`, or `0 ≤ y, z`, see `Ereal.left_distrib_of_nonneg` -(similarily for right distributivity). +(similarly for right distributivity). `EReal` is a `CompleteLinearOrder`; this is deduced by type class inference from the fact that `WithBot (WithTop L)` is a complete linear order if `L` is diff --git a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean index 02cc7fb943f06..a8d5a85385e1b 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean @@ -32,7 +32,7 @@ polynomial in `k` splits. algebraic closure, algebraically closed -## Main reults +## Main results - `IsAlgClosure.of_splits`: if `K / k` is algebraic, and every monic irreducible polynomial over `k` splits in `K`, then `K` is algebraically closed (in fact an algebraic closure of `k`). diff --git a/Mathlib/Geometry/Manifold/Instances/Sphere.lean b/Mathlib/Geometry/Manifold/Instances/Sphere.lean index abc2d5450eb3b..9ab63860e4d2a 100644 --- a/Mathlib/Geometry/Manifold/Instances/Sphere.lean +++ b/Mathlib/Geometry/Manifold/Instances/Sphere.lean @@ -203,7 +203,7 @@ theorem stereo_left_inv (hv : ‖v‖ = 1) {x : sphere (0 : E) 1} (hx : (x : E) stereoInvFun hv (stereoToFun v x) = x := by ext simp only [stereoToFun_apply, stereoInvFun_apply, smul_add] - -- name two frequently-occuring quantities and write down their basic properties + -- name two frequently-occurring quantities and write down their basic properties set a : ℝ := innerSL _ v x set y := orthogonalProjection (ℝ ∙ v)ᗮ x have split : ↑x = a • v + ↑y := by diff --git a/Mathlib/GroupTheory/Nilpotent.lean b/Mathlib/GroupTheory/Nilpotent.lean index 66235c6b0d78c..2b1d6262d8668 100644 --- a/Mathlib/GroupTheory/Nilpotent.lean +++ b/Mathlib/GroupTheory/Nilpotent.lean @@ -238,7 +238,7 @@ theorem nilpotent_iff_finite_ascending_central_series : rw [eq_top_iff, ← hn] exact ascending_central_series_le_upper H hH n -theorem is_decending_rev_series_of_is_ascending {H : ℕ → Subgroup G} {n : ℕ} (hn : H n = ⊤) +theorem is_descending_rev_series_of_is_ascending {H : ℕ → Subgroup G} {n : ℕ} (hn : H n = ⊤) (hasc : IsAscendingCentralSeries H) : IsDescendingCentralSeries fun m : ℕ => H (n - m) := by cases' hasc with h0 hH refine ⟨hn, fun x m hx g => ?_⟩ @@ -253,6 +253,9 @@ theorem is_decending_rev_series_of_is_ascending {H : ℕ → Subgroup G} {n : convert hx using 1 rw [tsub_add_eq_add_tsub (Nat.succ_le_of_lt hm), Nat.succ_eq_add_one, Nat.add_sub_add_right] +@[deprecated (since := "2024-12-25")] +alias is_decending_rev_series_of_is_ascending := is_descending_rev_series_of_is_ascending + theorem is_ascending_rev_series_of_is_descending {H : ℕ → Subgroup G} {n : ℕ} (hn : H n = ⊥) (hdesc : IsDescendingCentralSeries H) : IsAscendingCentralSeries fun m : ℕ => H (n - m) := by cases' hdesc with h0 hH @@ -273,7 +276,7 @@ theorem nilpotent_iff_finite_descending_central_series : rw [nilpotent_iff_finite_ascending_central_series] constructor · rintro ⟨n, H, hH, hn⟩ - refine ⟨n, fun m => H (n - m), is_decending_rev_series_of_is_ascending G hn hH, ?_⟩ + refine ⟨n, fun m => H (n - m), is_descending_rev_series_of_is_ascending G hn hH, ?_⟩ dsimp only rw [tsub_self] exact hH.1 @@ -396,7 +399,7 @@ theorem least_descending_central_series_length_eq_nilpotencyClass : rw [← least_ascending_central_series_length_eq_nilpotencyClass] refine le_antisymm (Nat.find_mono ?_) (Nat.find_mono ?_) · rintro n ⟨H, ⟨hH, hn⟩⟩ - refine ⟨fun m => H (n - m), is_decending_rev_series_of_is_ascending G hn hH, ?_⟩ + refine ⟨fun m => H (n - m), is_descending_rev_series_of_is_ascending G hn hH, ?_⟩ dsimp only rw [tsub_self] exact hH.1 diff --git a/Mathlib/GroupTheory/Sylow.lean b/Mathlib/GroupTheory/Sylow.lean index 8c1547040b62a..7c7d7efa40cf2 100644 --- a/Mathlib/GroupTheory/Sylow.lean +++ b/Mathlib/GroupTheory/Sylow.lean @@ -408,7 +408,7 @@ theorem card_dvd_index [Fact p.Prime] [Finite (Sylow p G)] (P : Sylow p G) : @[deprecated (since := "2024-11-07")] alias _root_.card_sylow_dvd_index := card_dvd_index -/-- Auxilliary lemma for `Sylow.not_dvd_index` which is strictly stronger. -/ +/-- Auxiliary lemma for `Sylow.not_dvd_index` which is strictly stronger. -/ private theorem not_dvd_index_aux [hp : Fact p.Prime] (P : Sylow p G) [P.Normal] [P.FiniteIndex] : ¬ p ∣ P.index := by intro h diff --git a/Mathlib/GroupTheory/Transfer.lean b/Mathlib/GroupTheory/Transfer.lean index 80b6e61edd5d6..bd55dae349bfd 100644 --- a/Mathlib/GroupTheory/Transfer.lean +++ b/Mathlib/GroupTheory/Transfer.lean @@ -306,7 +306,7 @@ namespace IsCyclic open Subgroup --- we could supress the variable `p`, but that might introduce `motive not type correct` issues. +-- we could suppress the variable `p`, but that might introduce `motive not type correct` issues. variable {G : Type*} [Group G] [Finite G] {p : ℕ} (hp : (Nat.card G).minFac = p) {P : Sylow p G} include hp in diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean index 259da107ba011..4cb74369937e0 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean @@ -710,7 +710,7 @@ section Ext variable [Finite ι] [DecidableEq ι] {f g : ((i : ι) → φv i) →ᵃ[k] P2} -/-- Two affine maps from a Pi-tyoe of modules `(i : ι) → φv i` are equal if they are equal in their +/-- Two affine maps from a Pi-type of modules `(i : ι) → φv i` are equal if they are equal in their operation on `Pi.single` and at zero. Analogous to `LinearMap.pi_ext`. See also `pi_ext_nonempty`, which instead of agreement at zero requires `Nonempty ι`. -/ theorem pi_ext_zero (h : ∀ i x, f (Pi.single i x) = g (Pi.single i x)) (h₂ : f 0 = g 0) : @@ -727,7 +727,7 @@ theorem pi_ext_zero (h : ∀ i x, f (Pi.single i x) = g (Pi.single i x)) (h₂ : rwa [s₂, s₃, h₂, vadd_right_cancel_iff] at s₁ · exact h₂ -/-- Two affine maps from a Pi-tyoe of modules `(i : ι) → φv i` are equal if they are equal in their +/-- Two affine maps from a Pi-type of modules `(i : ι) → φv i` are equal if they are equal in their operation on `Pi.single` and `ι` is nonempty. Analogous to `LinearMap.pi_ext`. See also `pi_ext_zero`, which instead `Nonempty ι` requires agreement at 0.-/ theorem pi_ext_nonempty [Nonempty ι] (h : ∀ i x, f (Pi.single i x) = g (Pi.single i x)) : diff --git a/Mathlib/LinearAlgebra/Eigenspace/Pi.lean b/Mathlib/LinearAlgebra/Eigenspace/Pi.lean index ddb8f9f9ea5ac..7e1f7f8a4f97e 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Pi.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Pi.lean @@ -10,7 +10,7 @@ import Mathlib.LinearAlgebra.Eigenspace.Triangularizable In finite dimensions, the theory of simultaneous eigenvalues for a family of linear endomorphisms `i ↦ f i` enjoys similar properties to that of a single endomorphism, provided the family obeys a -compatibilty condition. This condition is that the maximum generalised eigenspaces of each +compatibility condition. This condition is that the maximum generalised eigenspaces of each endomorphism are invariant under the action of all members of the family. It is trivially satisfied for commuting endomorphisms but there are important more general situations where it also holds (e.g., representations of nilpotent Lie algebras). diff --git a/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean b/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean index 3c18c28e836bd..5bd189f1e2b73 100644 --- a/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean +++ b/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean @@ -13,7 +13,7 @@ This file proves properties of 2×2 block matrices `[A B; C D]` that relate to t `D - C*A⁻¹*B`. Some of the results here generalize to 2×2 matrices in a category, rather than just a ring. A few -results in this direction can be found in the file `CateogryTheory.Preadditive.Biproducts`, +results in this direction can be found in the file `CategoryTheory.Preadditive.Biproducts`, especially the declarations `CategoryTheory.Biprod.gaussian` and `CategoryTheory.Biprod.isoElim`. Compare with `Matrix.invertibleOfFromBlocks₁₁Invertible`. diff --git a/Mathlib/LinearAlgebra/PerfectPairing.lean b/Mathlib/LinearAlgebra/PerfectPairing.lean index 78fff41b7e142..4227551192718 100644 --- a/Mathlib/LinearAlgebra/PerfectPairing.lean +++ b/Mathlib/LinearAlgebra/PerfectPairing.lean @@ -293,7 +293,7 @@ variable {S : Type*} {M' N' : Type*} [AddCommGroup M'] [Module S M'] [AddCommGroup N'] [Module S N'] (i : M' →ₗ[S] M) (j : N' →ₗ[S] N) -/-- An auxiliary definition used to constuct `PerfectPairing.restrictScalars`. -/ +/-- An auxiliary definition used to construct `PerfectPairing.restrictScalars`. -/ private def restrictScalarsAux (hp : ∀ m n, p (i m) (j n) ∈ (algebraMap S R).range) : M' →ₗ[S] N' →ₗ[S] S := diff --git a/Mathlib/MeasureTheory/Covering/Vitali.lean b/Mathlib/MeasureTheory/Covering/Vitali.lean index fc7dbaf32794b..416a61a4c48c3 100644 --- a/Mathlib/MeasureTheory/Covering/Vitali.lean +++ b/Mathlib/MeasureTheory/Covering/Vitali.lean @@ -15,9 +15,9 @@ Consider a family of balls `(B (x_i, r_i))_{i ∈ I}` in a metric space, with un radii. Then one can extract a disjoint subfamily indexed by `J ⊆ I`, such that any `B (x_i, r_i)` is included in a ball `B (x_j, 5 r_j)`. -We prove this theorem in `Vitali.exists_disjoint_subfamily_covering_enlargment_closedBall`. +We prove this theorem in `Vitali.exists_disjoint_subfamily_covering_enlargement_closedBall`. It is deduced from a more general version, called -`Vitali.exists_disjoint_subfamily_covering_enlargment`, which applies to any family of sets +`Vitali.exists_disjoint_subfamily_covering_enlargement`, which applies to any family of sets together with a size function `δ` (think "radius" or "diameter"). We deduce the measurable Vitali covering theorem. Assume one is given a family `t` of closed sets @@ -42,10 +42,10 @@ open scoped NNReal ENNReal Topology namespace Vitali /-- **Vitali covering theorem**: given a set `t` of subsets of a type, one may extract a disjoint -subfamily `u` such that the `τ`-enlargment of this family covers all elements of `t`, where `τ > 1` +subfamily `u` such that the `τ`-enlargement of this family covers all elements of `t`, where `τ > 1` is any fixed number. -When `t` is a family of balls, the `τ`-enlargment of `ball x r` is `ball x ((1+2τ) r)`. In general, +When `t` is a family of balls, the `τ`-enlargement of `ball x r` is `ball x ((1+2τ) r)`. In general, it is expressed in terms of a function `δ` (think "radius" or "diameter"), positive and bounded on all elements of `t`. The condition is that every element `a` of `t` should intersect an element `b` of `u` of size larger than that of `a` up to `τ`, i.e., `δ b ≥ δ a / τ`. @@ -53,7 +53,7 @@ element `b` of `u` of size larger than that of `a` up to `τ`, i.e., `δ b ≥ We state the lemma slightly more generally, with an indexed family of sets `B a` for `a ∈ t`, for wider applicability. -/ -theorem exists_disjoint_subfamily_covering_enlargment (B : ι → Set α) (t : Set ι) (δ : ι → ℝ) +theorem exists_disjoint_subfamily_covering_enlargement (B : ι → Set α) (t : Set ι) (δ : ι → ℝ) (τ : ℝ) (hτ : 1 < τ) (δnonneg : ∀ a ∈ t, 0 ≤ δ a) (R : ℝ) (δle : ∀ a ∈ t, δ a ≤ R) (hne : ∀ a ∈ t, (B a).Nonempty) : ∃ u ⊆ t, @@ -150,10 +150,14 @@ theorem exists_disjoint_subfamily_covering_enlargment (B : ι → Set α) (t : S · rw [← not_disjoint_iff_nonempty_inter] at hcb exact (hcb (H _ H')).elim +@[deprecated (since := "2024-12-25")] +alias exists_disjoint_subfamily_covering_enlargment := + exists_disjoint_subfamily_covering_enlargement + /-- Vitali covering theorem, closed balls version: given a family `t` of closed balls, one can extract a disjoint subfamily `u ⊆ t` so that all balls in `t` are covered by the τ-times dilations of balls in `u`, for some `τ > 3`. -/ -theorem exists_disjoint_subfamily_covering_enlargment_closedBall +theorem exists_disjoint_subfamily_covering_enlargement_closedBall [PseudoMetricSpace α] (t : Set ι) (x : ι → α) (r : ι → ℝ) (R : ℝ) (hr : ∀ a ∈ t, r a ≤ R) (τ : ℝ) (hτ : 3 < τ) : ∃ u ⊆ t, @@ -172,7 +176,7 @@ theorem exists_disjoint_subfamily_covering_enlargment_closedBall fun a ha => ⟨a, ha, by simp only [closedBall_eq_empty.2 (ht a ha), empty_subset]⟩⟩ push_neg at ht let t' := { a ∈ t | 0 ≤ r a } - rcases exists_disjoint_subfamily_covering_enlargment (fun a => closedBall (x a) (r a)) t' r + rcases exists_disjoint_subfamily_covering_enlargement (fun a => closedBall (x a) (r a)) t' r ((τ - 1) / 2) (by linarith) (fun a ha => ha.2) R (fun a ha => hr a ha.1) fun a ha => ⟨x a, mem_closedBall_self ha.2⟩ with ⟨u, ut', u_disj, hu⟩ @@ -190,6 +194,10 @@ theorem exists_disjoint_subfamily_covering_enlargment_closedBall rcases A b ⟨rb.1, rb.2⟩ with ⟨c, cu, _⟩ exact ⟨c, cu, by simp only [closedBall_eq_empty.2 h'a, empty_subset]⟩ +@[deprecated (since := "2024-12-25")] +alias exists_disjoint_subfamily_covering_enlargment_closedBall := + exists_disjoint_subfamily_covering_enlargement_closedBall + /-- The measurable Vitali covering theorem. Assume one is given a family `t` of closed sets with nonempty interior, such that each `a ∈ t` is included in a ball `B (x, r)` and covers a definite proportion of the ball `B (x, 3 r)` for a given measure `μ` (think of the situation where `μ` is @@ -243,7 +251,8 @@ theorem exists_disjoint_covering_ae apply ha.2.trans (hR1 (c a)) have A' : ∀ a ∈ t', (B a).Nonempty := fun a hat' => Set.Nonempty.mono interior_subset (ht a hat'.1) - refine exists_disjoint_subfamily_covering_enlargment B t' r 2 one_lt_two (fun a ha => ?_) 1 A A' + refine exists_disjoint_subfamily_covering_enlargement + B t' r 2 one_lt_two (fun a ha => ?_) 1 A A' exact nonempty_closedBall.1 ((A' a ha).mono (hB a ha.1)) have ut : u ⊆ t := fun a hau => (ut' hau).1 -- As the space is second countable, the family is countable since all its sets have nonempty diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index a7710637dad96..8633fe486564e 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -1492,7 +1492,7 @@ theorem measure_le_lintegral_thickenedIndicator (μ : Measure X) {E : Set X} end thickenedIndicator --- We declare a new `{X : Type*}` to discard the instance `[MeasureableSpace X]` +-- We declare a new `{X : Type*}` to discard the instance `[MeasurableSpace X]` -- which has been in scope for the entire file up to this point. variable {X : Type*} diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index 27a0544eab79a..5d1379bf8604a 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -21,7 +21,7 @@ the space such that the measurable space is generated by the union of all partit generated. * `MeasurableSpace.countableGeneratingSet`: a countable set of sets that generates the σ-algebra. * `MeasurableSpace.countablePartition`: sequences of finer and finer partitions of - a countably generated space, defined by taking the `memPartion` of an enumeration of the sets in + a countably generated space, defined by taking the `memPartition` of an enumeration of the sets in `countableGeneratingSet`. * `MeasurableSpace.SeparatesPoints` : class stating that a measurable space separates points. diff --git a/Mathlib/NumberTheory/Cyclotomic/Three.lean b/Mathlib/NumberTheory/Cyclotomic/Three.lean index a022fd1146b1b..e82393c894547 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Three.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Three.lean @@ -128,7 +128,7 @@ lemma lambda_dvd_or_dvd_sub_one_or_dvd_add_one [NumberField K] [IsCyclotomicExte have := hζ.finite_quotient_toInteger_sub_one (by decide) let _ := Fintype.ofFinite (𝓞 K ⧸ Ideal.span {λ}) let _ : Ring (𝓞 K ⧸ Ideal.span {λ}) := CommRing.toRing -- to speed up instance synthesis - let _ : AddGroup (𝓞 K ⧸ Ideal.span {λ}) := AddGroupWithOne.toAddGroup -- dito + let _ : AddGroup (𝓞 K ⧸ Ideal.span {λ}) := AddGroupWithOne.toAddGroup -- ditto have := Finset.mem_univ (Ideal.Quotient.mk (Ideal.span {λ}) x) have h3 : Fintype.card (𝓞 K ⧸ Ideal.span {λ}) = 3 := by rw [← Nat.card_eq_fintype_card, hζ.card_quotient_toInteger_sub_one (by decide), diff --git a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean index 840e9ccbfa672..7139b4bb1fc27 100644 --- a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean +++ b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean @@ -217,7 +217,7 @@ private lemma summable_F'' : Summable F'' := by show (1 : ℝ) - 2⁻¹ = 2⁻¹ by norm_num, inv_le_inv₀ (mod_cast p.prop.pos) zero_lt_two] exact Nat.ofNat_le_cast.mpr p.prop.two_le -/-- The function `n ↦ Λ n / n`, restriced to non-primes in a residue class, is summable. +/-- The function `n ↦ Λ n / n`, restricted to non-primes in a residue class, is summable. This is used to convert results on `ArithmeticFunction.vonMangoldt.residueClass` to results on primes in an arithmetic progression. -/ lemma summable_residueClass_non_primes_div : @@ -257,7 +257,7 @@ lemma summable_residueClass_non_primes_div : variable [NeZero q] {a} /-- We can express `ArithmeticFunction.vonMangoldt.residueClass` as a linear combination -of twists of the von Mangoldt function by Dirichlet charaters. -/ +of twists of the von Mangoldt function by Dirichlet characters. -/ lemma residueClass_apply (ha : IsUnit a) (n : ℕ) : residueClass a n = (q.totient : ℂ)⁻¹ * ∑ χ : DirichletCharacter ℂ q, χ a⁻¹ * χ n * vonMangoldt n := by @@ -267,7 +267,7 @@ lemma residueClass_apply (ha : IsUnit a) (n : ℕ) : ite_mul, zero_mul, ↓reduceIte, ite_self] /-- We can express `ArithmeticFunction.vonMangoldt.residueClass` as a linear combination -of twists of the von Mangoldt function by Dirichlet charaters. -/ +of twists of the von Mangoldt function by Dirichlet characters. -/ lemma residueClass_eq (ha : IsUnit a) : ↗(residueClass a) = (q.totient : ℂ)⁻¹ • ∑ χ : DirichletCharacter ℂ q, χ a⁻¹ • (fun n : ℕ ↦ χ n * vonMangoldt n) := by diff --git a/Mathlib/NumberTheory/NumberField/Completion.lean b/Mathlib/NumberTheory/NumberField/Completion.lean index 0d23896da11ed..9a37fd57f1c6b 100644 --- a/Mathlib/NumberTheory/NumberField/Completion.lean +++ b/Mathlib/NumberTheory/NumberField/Completion.lean @@ -118,7 +118,7 @@ theorem extensionEmbedding_of_isReal_coe {v : InfinitePlace K} (hv : IsReal v) ( theorem isometry_extensionEmbedding : Isometry (extensionEmbedding v) := Isometry.of_dist_eq (extensionEmbedding_dist_eq_of_comp v.norm_embedding_eq) -/-- The embedding `v.Completion →+* ℝ` at a real infinite palce is an isometry. -/ +/-- The embedding `v.Completion →+* ℝ` at a real infinite place is an isometry. -/ theorem isometry_extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : Isometry (extensionEmbeddingOfIsReal hv) := Isometry.of_dist_eq (extensionEmbedding_dist_eq_of_comp <| v.norm_embedding_of_isReal hv) diff --git a/Mathlib/NumberTheory/Padics/MahlerBasis.lean b/Mathlib/NumberTheory/Padics/MahlerBasis.lean index 6f4ac315c6928..ad170376bf16d 100644 --- a/Mathlib/NumberTheory/Padics/MahlerBasis.lean +++ b/Mathlib/NumberTheory/Padics/MahlerBasis.lean @@ -326,7 +326,7 @@ variable {p : ℕ} [hp : Fact p.Prime] {E : Type*} /-- **Mahler's theorem**: for any continuous function `f` from `ℤ_[p]` to a `p`-adic Banach space, the -Mahler series with coeffients `n ↦ Δ_[1]^[n] f 0` converges to the original function `f`. +Mahler series with coefficients `n ↦ Δ_[1]^[n] f 0` converges to the original function `f`. -/ lemma hasSum_mahler (f : C(ℤ_[p], E)) : HasSum (fun n ↦ mahlerTerm (Δ_[1]^[n] f 0) n) f := by -- First show `∑' n, mahler_term f n` converges to *something*. diff --git a/Mathlib/NumberTheory/RamificationInertia/Basic.lean b/Mathlib/NumberTheory/RamificationInertia/Basic.lean index d0291b1916d6f..24f96fd9081fb 100644 --- a/Mathlib/NumberTheory/RamificationInertia/Basic.lean +++ b/Mathlib/NumberTheory/RamificationInertia/Basic.lean @@ -465,7 +465,7 @@ theorem Quotient.algebraMap_quotient_pow_ramificationIdx (x : R) : /-- 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. +This can't be an instance since the map `f : R → S` is generally not inferable. -/ def Quotient.algebraQuotientOfRamificationIdxNeZero [hfp : NeZero e] : Algebra (R ⧸ p) (S ⧸ P) := diff --git a/Mathlib/Order/Grade.lean b/Mathlib/Order/Grade.lean index 9818a18268e0b..72f89b3e3df10 100644 --- a/Mathlib/Order/Grade.lean +++ b/Mathlib/Order/Grade.lean @@ -296,13 +296,13 @@ abbrev GradeBoundedOrder.liftRight [GradeBoundedOrder 𝕆 β] (f : α → β) ( -- See note [reducible non-instances] /-- A `Fin n`-graded order is also `ℕ`-graded. We do not mark this an instance because `n` is not -inferrable. -/ +inferable. -/ abbrev GradeOrder.finToNat (n : ℕ) [GradeOrder (Fin n) α] : GradeOrder ℕ α := (GradeOrder.liftLeft (_ : Fin n → ℕ) Fin.val_strictMono) fun _ _ => CovBy.coe_fin -- See note [reducible non-instances] /-- A `Fin n`-graded order is also `ℕ`-graded. We do not mark this an instance because `n` is not -inferrable. -/ +inferable. -/ abbrev GradeMinOrder.finToNat (n : ℕ) [GradeMinOrder (Fin n) α] : GradeMinOrder ℕ α := (GradeMinOrder.liftLeft (_ : Fin n → ℕ) Fin.val_strictMono fun _ _ => CovBy.coe_fin) fun a h => by cases n diff --git a/Mathlib/Order/KrullDimension.lean b/Mathlib/Order/KrullDimension.lean index a6ba1d3f81334..3a97b0b3185a7 100644 --- a/Mathlib/Order/KrullDimension.lean +++ b/Mathlib/Order/KrullDimension.lean @@ -474,7 +474,7 @@ lemma coheight_eq_coe_iff {x : α} {n : ℕ} : coheight x < ⊤ ∧ (n = 0 ∨ ∃ y > x, coheight y = n - 1) ∧ (∀ y > x, coheight y < n) := height_eq_coe_iff (α := αᵒᵈ) -/-- The elements of finite height `n` are the minimial elements among those of height `≥ n`. -/ +/-- The elements of finite height `n` are the minimal elements among those of height `≥ n`. -/ lemma height_eq_coe_iff_minimal_le_height {a : α} {n : ℕ} : height a = n ↔ Minimal (fun y => n ≤ height y) a := by by_cases hfin : height a < ⊤ diff --git a/Mathlib/Order/SuccPred/WithBot.lean b/Mathlib/Order/SuccPred/WithBot.lean index 01442b1093c2f..cac81174a0073 100644 --- a/Mathlib/Order/SuccPred/WithBot.lean +++ b/Mathlib/Order/SuccPred/WithBot.lean @@ -44,7 +44,7 @@ end WithBot namespace WithTop variable {α : Type*} [Preorder α] [OrderTop α] [PredOrder α] {x y : WithTop α} -/-- The predessor of `a : WithTop α` as an element of `α`. -/ +/-- The predecessor of `a : WithTop α` as an element of `α`. -/ def pred (a : WithTop α) : α := a.recTopCoe ⊤ Order.pred /-- Not to be confused with `WithTop.orderPred_top`, which is about `Order.pred`. -/ diff --git a/Mathlib/Probability/Independence/Kernel.lean b/Mathlib/Probability/Independence/Kernel.lean index 27dcf76e54543..c83e793c26c67 100644 --- a/Mathlib/Probability/Independence/Kernel.lean +++ b/Mathlib/Probability/Independence/Kernel.lean @@ -18,7 +18,7 @@ for any sets `f i_1 ∈ π i_1, ..., f i_n ∈ π i_n`, then for `μ`-almost eve This notion of independence is a generalization of both independence and conditional independence. For conditional independence, `κ` is the conditional kernel `ProbabilityTheory.condexpKernel` and -`μ` is the ambiant measure. For (non-conditional) independence, `κ = Kernel.const Unit μ` and the +`μ` is the ambient measure. For (non-conditional) independence, `κ = Kernel.const Unit μ` and the measure is the Dirac measure on `Unit`. The main purpose of this file is to prove only once the properties that hold for both conditional diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index 52b8846b37708..d07c60a4fa0bb 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -439,7 +439,7 @@ theorem val_natCast (n : ℕ) : (n : HomogeneousLocalization 𝒜 x).val = n := theorem val_intCast (n : ℤ) : (n : HomogeneousLocalization 𝒜 x).val = n := show val (Int.castDef n) = _ by cases n <;> simp [Int.castDef, *] -instance homogenousLocalizationCommRing : CommRing (HomogeneousLocalization 𝒜 x) := +instance homogeneousLocalizationCommRing : CommRing (HomogeneousLocalization 𝒜 x) := (HomogeneousLocalization.val_injective x).commRing _ val_zero val_one val_add val_mul val_neg val_sub (val_nsmul x · ·) (val_zsmul x · ·) val_pow val_natCast val_intCast diff --git a/Mathlib/RingTheory/Lasker.lean b/Mathlib/RingTheory/Lasker.lean index 93d088c47875b..fc3785bfee908 100644 --- a/Mathlib/RingTheory/Lasker.lean +++ b/Mathlib/RingTheory/Lasker.lean @@ -15,9 +15,9 @@ import Mathlib.RingTheory.Noetherian.Defs - `IsLasker`: A ring `R` satisfies `IsLasker R` when any `I : Ideal R` can be decomposed into finitely many primary ideals. -- `IsLasker.minimal`: Any `I : Ideal R` in a ring `R` satisifying `IsLasker R` can be +- `IsLasker.minimal`: Any `I : Ideal R` in a ring `R` satisfying `IsLasker R` can be decomposed into primary ideals, such that the decomposition is minimal: - each primary ideal is necessary, and each primary ideal has an indepedent radical. + each primary ideal is necessary, and each primary ideal has an independent radical. - `Ideal.isLasker`: Every Noetherian commutative ring is a Lasker ring. ## Implementation details diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index c66af8206d58a..454a359a1e689 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -51,7 +51,7 @@ series to which the filter `ℱ` converges. * The field of rational functions is dense in `LaurentSeries`: this is the declaration `LaurentSeries.coe_range_dense` and relies principally upon `LaurentSeries.exists_ratFunc_val_lt`, stating that for every Laurent series `f` and every `γ : ℤₘ₀` one can find a rational function `Q` -such that the `X`-adic valuation `v` satifies `v (f - Q) < γ`. +such that the `X`-adic valuation `v` satisfies `v (f - Q) < γ`. * In `LaurentSeries.valuation_compare` we prove that the extension of the `X`-adic valuation from `RatFunc K` up to its abstract completion coincides, modulo the isomorphism with `K⸨X⸩`, with the `X`-adic valuation on `K⸨X⸩`. @@ -918,7 +918,7 @@ theorem exists_Polynomial_intValuation_lt (F : K⟦X⟧) (η : ℤₘ₀ˣ) : exact Int.zero_lt_one /-- For every Laurent series `f` and every `γ : ℤₘ₀` one can find a rational function `Q` such -that the `X`-adic valuation `v` satifies `v (f - Q) < γ`. -/ +that the `X`-adic valuation `v` satisfies `v (f - Q) < γ`. -/ theorem exists_ratFunc_val_lt (f : K⸨X⸩) (γ : ℤₘ₀ˣ) : ∃ Q : RatFunc K, Valued.v (f - Q) < γ := by set F := f.powerSeriesPart with hF diff --git a/Mathlib/RingTheory/MvPolynomial/EulerIdentity.lean b/Mathlib/RingTheory/MvPolynomial/EulerIdentity.lean index 80d599cc70bfb..30c453aa65390 100644 --- a/Mathlib/RingTheory/MvPolynomial/EulerIdentity.lean +++ b/Mathlib/RingTheory/MvPolynomial/EulerIdentity.lean @@ -9,7 +9,7 @@ import Mathlib.RingTheory.MvPolynomial.Homogeneous /-! # Euler's homogeneous identity -## Main resutls +## Main results * `IsHomogeneous.sum_X_mul_pderiv`: Euler's identity for homogeneous polynomials: for a multivariate homogeneous polynomial, diff --git a/Mathlib/RingTheory/MvPowerSeries/Order.lean b/Mathlib/RingTheory/MvPowerSeries/Order.lean index 2a18f68077597..2b960ab4d11a8 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Order.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Order.lean @@ -33,7 +33,7 @@ than the weighted order vanish - `MvPowerSeries.weightedOrder_eq_top_iff`: the weighted order of `f` is `⊤` if and only `f = 0`. -- `MvPowerSerie.nat_le_weightedOrder`: if all coefficients of weight `< n` vanish, then the +- `MvPowerSeries.nat_le_weightedOrder`: if all coefficients of weight `< n` vanish, then the weighted order is at least `n`. - `MvPowerSeries.weightedOrder_eq_nat_iff`: the weighted order is some integer `n` iff there diff --git a/Mathlib/RingTheory/PowerSeries/PiTopology.lean b/Mathlib/RingTheory/PowerSeries/PiTopology.lean index b3cf011bd7091..6c12ec4d04fac 100644 --- a/Mathlib/RingTheory/PowerSeries/PiTopology.lean +++ b/Mathlib/RingTheory/PowerSeries/PiTopology.lean @@ -31,7 +31,7 @@ of `f` is nilpotent, or vanishes, then the powers of `f` converge to zero. converge to zero iff the constant coefficient of `f` is nilpotent. - `PowerSeries.WithPiTopology.hasSum_of_monomials_self` : viewed as an infinite sum, a power -series coverges to itself. +series converges to itself. TODO: add the similar result for the series of homogeneous components. diff --git a/Mathlib/RingTheory/RootsOfUnity/AlgebraicallyClosed.lean b/Mathlib/RingTheory/RootsOfUnity/AlgebraicallyClosed.lean index ce7dadf0451bc..4c6fb41681723 100644 --- a/Mathlib/RingTheory/RootsOfUnity/AlgebraicallyClosed.lean +++ b/Mathlib/RingTheory/RootsOfUnity/AlgebraicallyClosed.lean @@ -10,7 +10,7 @@ import Mathlib.NumberTheory.Cyclotomic.Basic # Instances for HasEnoughRootsOfUnity We provide an instance for `HasEnoughRootsOfUnity F n` when `F` is an algebraically closed field -and `n` is not divisible by the characteristic. In particular, when `F` has characterstic zero, +and `n` is not divisible by the characteristic. In particular, when `F` has characteristic zero, this hold for all `n ≠ 0`. ### TODO diff --git a/Mathlib/RingTheory/Valuation/ValuationSubring.lean b/Mathlib/RingTheory/Valuation/ValuationSubring.lean index 813f07bdbde31..8f4342e738c9a 100644 --- a/Mathlib/RingTheory/Valuation/ValuationSubring.lean +++ b/Mathlib/RingTheory/Valuation/ValuationSubring.lean @@ -32,7 +32,7 @@ variable (K : Type u) [Field K] /-- A valuation subring of a field `K` is a subring `A` such that for every `x : K`, either `x ∈ A` or `x⁻¹ ∈ A`. -This is equvalent to being maximal in the domination order +This is equivalent to being maximal in the domination order of local subrings (the stacks project definition). See `LocalSubring.isMax_iff`. -/ structure ValuationSubring extends Subring K where diff --git a/Mathlib/Tactic/Algebraize.lean b/Mathlib/Tactic/Algebraize.lean index 237bc5f54de7a..1a6a8f3e8f549 100644 --- a/Mathlib/Tactic/Algebraize.lean +++ b/Mathlib/Tactic/Algebraize.lean @@ -38,7 +38,7 @@ specified declaration should be one of the following: 1. An inductive type (i.e. the `Algebra` property itself), in this case it is assumed that the `RingHom` and the `Algebra` property are definitionally the same, and the tactic will construct the -`Algebra` property by giving the `RingHom` property as a term. Due to how this is peformed, we also +`Algebra` property by giving the `RingHom` property as a term. Due to how this is performed, we also need to assume that the `Algebra` property can be constructed only from the homomorphism, so it can not have any other explicit arguments. 2. A lemma (or constructor) proving the `Algebra` property from the `RingHom` property. In this case @@ -124,7 +124,7 @@ namespace Algebraize /-- Given an expression `f` of type `RingHom A B` where `A` and `B` are commutative semirings, this function adds the instance `Algebra A B` to the context (if it does not already exist). -This function also requries the type of `f`, given by the parameter `ft`. The reason this is done +This function also requires the type of `f`, given by the parameter `ft`. The reason this is done (even though `ft` can be inferred from `f`) is to avoid recomputing `ft` in the `algebraize` tactic, as when `algebraize` calls `addAlgebraInstanceFromRingHom` it has already computed `ft`. -/ def addAlgebraInstanceFromRingHom (f ft : Expr) : TacticM Unit := withMainContext do diff --git a/Mathlib/Tactic/FunProp/Core.lean b/Mathlib/Tactic/FunProp/Core.lean index 105bd3beea9c0..e5b962482c3de 100644 --- a/Mathlib/Tactic/FunProp/Core.lean +++ b/Mathlib/Tactic/FunProp/Core.lean @@ -625,7 +625,7 @@ mutual withTraceNode `Meta.Tactic.fun_prop (fun r => do pure s!"[{ExceptToEmoji.toEmoji r}] {← ppExpr e}") do - -- check cache for succesfull goals + -- check cache for successful goals if let .some { expr := _, proof? := .some proof } := (← get).cache.find? e then trace[Meta.Tactic.fun_prop] "reusing previously found proof for {e}" return .some { proof := proof } diff --git a/Mathlib/Tactic/FunProp/Theorems.lean b/Mathlib/Tactic/FunProp/Theorems.lean index 464ce021ba426..0ed63b6296c34 100644 --- a/Mathlib/Tactic/FunProp/Theorems.lean +++ b/Mathlib/Tactic/FunProp/Theorems.lean @@ -268,7 +268,7 @@ initialize morTheoremsExt : GeneralTheoremsExt ← /-- Get morphism theorems applicable to `e`. For example calling on `e` equal to `Continuous f` for `f : X→L[ℝ] Y` would return theorem -infering continuity from the bundled morphism. -/ +inferring continuity from the bundled morphism. -/ def getMorphismTheorems (e : Expr) : FunPropM (Array GeneralTheorem) := do let ext := morTheoremsExt.getState (← getEnv) let candidates ← withConfig (fun cfg => { cfg with iota := false, zeta := false }) <| diff --git a/Mathlib/Tactic/ReduceModChar.lean b/Mathlib/Tactic/ReduceModChar.lean index ed98906d7949f..18b8a08cf04ce 100644 --- a/Mathlib/Tactic/ReduceModChar.lean +++ b/Mathlib/Tactic/ReduceModChar.lean @@ -82,7 +82,7 @@ mutual assumeInstancesCommute return .isNat sα c q(CharP.isNat_pow (f := $f) $instCharP (.refl $f) $pa $pb $pn $r) - /-- If `e` is of the form `a ^ b`, reduce it using fast modular exponentation, otherwise + /-- If `e` is of the form `a ^ b`, reduce it using fast modular exponentiation, otherwise reduce it using `norm_num`. -/ partial def normIntNumeral' {α : Q(Type u)} (n n' : Q(ℕ)) (pn : Q(IsNat «$n» «$n'»)) (e : Q($α)) (_ : Q(Ring $α)) (instCharP : Q(CharP $α $n)) : MetaM (Result e) := diff --git a/Mathlib/Tactic/UnsetOption.lean b/Mathlib/Tactic/UnsetOption.lean index f91917e409006..6d9929d02792a 100644 --- a/Mathlib/Tactic/UnsetOption.lean +++ b/Mathlib/Tactic/UnsetOption.lean @@ -13,7 +13,7 @@ import Lean.Elab.Command This file defines an `unset_option` user command, which unsets user configurable options. -For example inputing `set_option blah 7` and then `unset_option blah` +For example, inputting `set_option blah 7` and then `unset_option blah` returns the user to the default state before any `set_option` command is called. This is helpful when the user does not know the default value of the option or it is cleaner not to write it explicitly, or for some options where the default diff --git a/Mathlib/Tactic/Variable.lean b/Mathlib/Tactic/Variable.lean index 3f885e361c071..01eccbf6ff905 100644 --- a/Mathlib/Tactic/Variable.lean +++ b/Mathlib/Tactic/Variable.lean @@ -72,7 +72,7 @@ structure VectorSpace (k V : Type*) [Field k] [AddCommGroup V] [Module k V] then `variable? [VectorSpace k V]` is equivalent to `variable {k V : Type*} [Field k] [AddCommGroup V] [Module k V]`, assuming that there are no pre-existing instances on `k` and `V`. -Note that this is not a simple replacement: it only adds instances not inferrable +Note that this is not a simple replacement: it only adds instances not inferable from others in the current scope. A word of warning: the core algorithm depends on pretty printing, so if terms that appear diff --git a/Mathlib/Topology/Connected/Clopen.lean b/Mathlib/Topology/Connected/Clopen.lean index 6b1fcb3fcbf35..adb445b03e6d9 100644 --- a/Mathlib/Topology/Connected/Clopen.lean +++ b/Mathlib/Topology/Connected/Clopen.lean @@ -17,7 +17,7 @@ to clopen sets. + `IsClopen.biUnion_connectedComponent_eq`: a clopen set is the union of its connected components. + `PreconnectedSpace.induction₂`: an induction principle for preconnected spaces. -+ `ConnectedComponents`: The connected compoenents of a topological space, as a quotient type. ++ `ConnectedComponents`: The connected components of a topological space, as a quotient type. -/ diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index 24523cb3973cc..94bdce684cd4a 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -12,7 +12,7 @@ In this file we introduce 3 typeclass mixins that relate topology and order stru - `ClosedIicTopology` says that all the intervals $(-∞, a]$ (formally, `Set.Iic a`) are closed sets; -- `ClosedIciTopoplogy` says that all the intervals $[a, +∞)$ (formally, `Set.Ici a`) +- `ClosedIciTopology` says that all the intervals $[a, +∞)$ (formally, `Set.Ici a`) are closed sets; - `OrderClosedTopology` says that the set of points `(x, y)` such that `x ≤ y` is closed in the product topology. @@ -25,7 +25,7 @@ We prove many basic properties of such topologies. This file contains the proofs of the following facts. For exact requirements -(`OrderClosedTopology` vs `ClosedIciTopoplogy` vs `ClosedIicTopology, +(`OrderClosedTopology` vs `ClosedIciTopology` vs `ClosedIicTopology, `Preorder` vs `PartialOrder` vs `LinearOrder` etc) see their statements. diff --git a/Mathlib/Topology/UniformSpace/Dini.lean b/Mathlib/Topology/UniformSpace/Dini.lean index 9e90b66608f20..aff1384ce1d4f 100644 --- a/Mathlib/Topology/UniformSpace/Dini.lean +++ b/Mathlib/Topology/UniformSpace/Dini.lean @@ -65,7 +65,7 @@ lemma tendstoLocallyUniformly_of_forall_tendsto exact hF_mono hnm z /-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions on a -set `s` converging pointwise to a continuous funciton `f`, then `F n` converges locally uniformly +set `s` converging pointwise to a continuous function `f`, then `F n` converges locally uniformly to `f`. -/ lemma tendstoLocallyUniformlyOn_of_forall_tendsto {s : Set α} (hF_cont : ∀ i, ContinuousOn (F i) s) (hF_mono : ∀ x ∈ s, Monotone (F · x)) diff --git a/MathlibTest/algebraize.lean b/MathlibTest/algebraize.lean index bdd3c99219a1a..43b6fab39d5df 100644 --- a/MathlibTest/algebraize.lean +++ b/MathlibTest/algebraize.lean @@ -15,25 +15,25 @@ see e.g. `RingHom.FiniteType` for a concrete example of this. -/ def RingHom.testProperty1 {A B : Type*} [CommRing A] [CommRing B] (f : A →+* B) : Prop := @Algebra.testProperty1 A B _ _ f.toAlgebra -/-- Test property for when the `RingHom` porperty corresponds to a `Module` property (that is +/-- Test property for when the `RingHom` property corresponds to a `Module` property (that is definitionally the same). See e.g. `Module.Finite` for a concrete example of this. -/ class Module.testProperty2 (A M : Type*) [Semiring A] [AddCommMonoid M] [Module A M] : Prop where out : ∀ x : A, ∀ M : M, x • M = 0 -/-- Test property for when the `RingHom` porperty corresponds to a `Module` property (that is +/-- Test property for when the `RingHom` property corresponds to a `Module` property (that is definitionally the same). See e.g. `Module.Finite` for a concrete example of this. -/ @[algebraize Module.testProperty2] def RingHom.testProperty2 {A B : Type*} [CommRing A] [CommRing B] (f : A →+* B) : Prop := letI : Algebra A B := f.toAlgebra Module.testProperty2 A B -/-- Test property for when the `RingHom` porperty corresponds to a `Algebra` property that is not +/-- Test property for when the `RingHom` property corresponds to a `Algebra` property that is not definitionally the same, and needs to be created through a lemma. See e.g. `Algebra.IsIntegral` for an example. -/ class Algebra.testProperty3 (A B : Type*) [CommRing A] [CommRing B] [Algebra A B] : Prop where out : Algebra.testProperty1 A B -/- Test property for when the `RingHom` porperty corresponds to a `Algebra` property that is not +/- Test property for when the `RingHom` property corresponds to a `Algebra` property that is not definitionally the same, and needs to be created through a lemma. See e.g. `Algebra.IsIntegral` for an example. -/ @[algebraize Algebra.testProperty3.mk] diff --git a/MathlibTest/zmod.lean b/MathlibTest/zmod.lean index 19fe7d3bb0c27..f33029e4fd736 100644 --- a/MathlibTest/zmod.lean +++ b/MathlibTest/zmod.lean @@ -8,7 +8,7 @@ def p := 2^k - 1 set_option exponentiation.threshold 100000 in -- We ensure here that the `@[csimp]` attribute successfully replaces (at runtime) --- the default value of `npowRec` for the exponentation operation in `Monoid` +-- the default value of `npowRec` for the exponentiation operation in `Monoid` -- with a tail-recursive implementation by repeated squaring. /-- info: 1 -/ #guard_msgs in diff --git a/scripts/add_deprecations.sh b/scripts/add_deprecations.sh index a607c54526826..ae89ebfeaf5d5 100755 --- a/scripts/add_deprecations.sh +++ b/scripts/add_deprecations.sh @@ -9,7 +9,7 @@ confused. ** Assumptions ** The script works under the assumption that the only modifications between `master` and the current -branch are renames of lemmas that sould be deprecated. +branch are renames of lemmas that should be deprecated. The script inspects the relevant `git diff`, extracts the old name and the new one and uses this information to insert deprecated aliases as needed. diff --git a/scripts/late_importers.sh b/scripts/late_importers.sh index 86d540c8afa72..36273a05a12ed 100755 --- a/scripts/late_importers.sh +++ b/scripts/late_importers.sh @@ -25,7 +25,7 @@ The script that takes 4 arguments: 1. the module name that should be built (`Mathlib` by default); 2. the number of module names of output that the script returns (all by default); 3. the threshold difference in imports above which the script reports a module (all by default); -4. the GitHub jod-ID of the workflow run. +4. the GitHub job-ID of the workflow run. TODO * If a merged PR touched one of the files in the most recent report, trigger the workflow again. From a0a62fedc7889f5be805784547267cf20f507790 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 26 Dec 2024 12:12:45 +0000 Subject: [PATCH 004/189] feat(Algebra/Colimit/ModuleRing): split into Module and Ring (#20247) 20 files see a 100+ decrease in transitive imports. --- Mathlib.lean | 3 +- .../Algebra/Category/AlgebraCat/Limits.lean | 1 + .../Algebra/Category/ModuleCat/Limits.lean | 2 +- Mathlib/Algebra/Colimit/DirectLimit.lean | 13 +- .../Colimit/{ModuleRing.lean => Module.lean} | 337 +---------------- Mathlib/Algebra/Colimit/Ring.lean | 356 ++++++++++++++++++ Mathlib/Algebra/Homology/LocalCohomology.lean | 1 + .../TensorProduct/DirectLimit.lean | 3 +- Mathlib/RingTheory/FreeCommRing.lean | 2 +- 9 files changed, 376 insertions(+), 342 deletions(-) rename Mathlib/Algebra/Colimit/{ModuleRing.lean => Module.lean} (55%) create mode 100644 Mathlib/Algebra/Colimit/Ring.lean diff --git a/Mathlib.lean b/Mathlib.lean index 25b6753d6ca64..c7c011d73f350 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -182,7 +182,8 @@ import Mathlib.Algebra.CharZero.Infinite import Mathlib.Algebra.CharZero.Lemmas import Mathlib.Algebra.CharZero.Quotient import Mathlib.Algebra.Colimit.DirectLimit -import Mathlib.Algebra.Colimit.ModuleRing +import Mathlib.Algebra.Colimit.Module +import Mathlib.Algebra.Colimit.Ring import Mathlib.Algebra.ContinuedFractions.Basic import Mathlib.Algebra.ContinuedFractions.Computation.ApproximationCorollaries import Mathlib.Algebra.ContinuedFractions.Computation.Approximations diff --git a/Mathlib/Algebra/Category/AlgebraCat/Limits.lean b/Mathlib/Algebra/Category/AlgebraCat/Limits.lean index ea39419a2442e..33cd9ab524d58 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Limits.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Limits.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ +import Mathlib.Algebra.Algebra.Pi import Mathlib.Algebra.Category.AlgebraCat.Basic import Mathlib.Algebra.Category.ModuleCat.Basic import Mathlib.Algebra.Category.ModuleCat.Limits diff --git a/Mathlib/Algebra/Category/ModuleCat/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Limits.lean index 0f97bf8276775..431ac7e1734a3 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Limits.lean @@ -5,7 +5,7 @@ Authors: Kim Morrison -/ import Mathlib.Algebra.Category.ModuleCat.Basic import Mathlib.Algebra.Category.Grp.Limits -import Mathlib.Algebra.Colimit.ModuleRing +import Mathlib.Algebra.Colimit.Module /-! # The category of R-modules has all limits diff --git a/Mathlib/Algebra/Colimit/DirectLimit.lean b/Mathlib/Algebra/Colimit/DirectLimit.lean index 3da4299d697b6..c71463fd88e85 100644 --- a/Mathlib/Algebra/Colimit/DirectLimit.lean +++ b/Mathlib/Algebra/Colimit/DirectLimit.lean @@ -21,18 +21,19 @@ from magma (`Mul`) to `Field`. To make everything "hom-polymorphic", we work wit of `FunLike`s rather than plain unbundled functions, and we use algebraic hom typeclasses (e.g. `LinearMapClass`, `RingHomClass`) everywhere. -In `Mathlib.Algebra.Colimit.ModuleRing`, `Module.DirectLimit`, `AddCommGroup.DirectLimit` and -`Ring.DirectLimit` are defined as quotients of the universal objects (`DirectSum` and -`FreeCommRing`). These definitions are more general and suitable for arbitrary colimits, -but do not immediately provide criteria to determine when two elements in a component are equal +In `Mathlib.Algebra.Colimit.Module` and `Mathlib.Algebra.Colimit.Ring`, +`Module.DirectLimit`, `AddCommGroup.DirectLimit` and `Ring.DirectLimit` +are defined as quotients of the universal objects (`DirectSum` and `FreeCommRing`). +These definitions are more general and suitable for arbitrary colimits, but do not +immediately provide criteria to determine when two elements in a component are equal in the direct limit. On the other hand, the `DirectLimit` in this file is only defined for directed systems and does not work for general colimits, but the equivalence relation defining `DirectLimit` is very explicit. For colimits of directed systems there is no need to construct the universal object for each type of algebraic structure; the same type `DirectLimit` simply -works for all of them. This file is therefore more general than the `ModuleRing` file in -terms of the variety of algebraic structures supported. +works for all of them. This file is therefore more general than the `Module` and `Ring` +files in terms of the variety of algebraic structures supported. So far we only show that `DirectLimit` is the colimit in the categories of modules and rings, but for the other algebraic structures the constructions and proofs will be easy following diff --git a/Mathlib/Algebra/Colimit/ModuleRing.lean b/Mathlib/Algebra/Colimit/Module.lean similarity index 55% rename from Mathlib/Algebra/Colimit/ModuleRing.lean rename to Mathlib/Algebra/Colimit/Module.lean index 5b66a75497403..075b0220f0694 100644 --- a/Mathlib/Algebra/Colimit/ModuleRing.lean +++ b/Mathlib/Algebra/Colimit/Module.lean @@ -3,31 +3,27 @@ Copyright (c) 2019 Kenny Lau, Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Chris Hughes, Jujian Zhang -/ +import Mathlib.Algebra.Colimit.DirectLimit import Mathlib.Algebra.DirectSum.Module import Mathlib.Data.Finset.Order -import Mathlib.RingTheory.FreeCommRing -import Mathlib.RingTheory.Ideal.Maps -import Mathlib.RingTheory.Ideal.Quotient.Defs -import Mathlib.Algebra.Colimit.DirectLimit +import Mathlib.GroupTheory.Congruence.Hom import Mathlib.Tactic.SuppressCompilation /-! -# Direct limit of modules, abelian groups, rings, and fields. +# Direct limit of modules and abelian groups See Atiyah-Macdonald PP.32-33, Matsumura PP.269-270 Generalizes the notion of "union", or "gluing", of incomparable modules over the same ring, -or incomparable abelian groups, or rings, or fields. +or incomparable abelian groups. -It is constructed as a quotient of the free module (for the module case) or quotient of -the free commutative ring (for the ring case) instead of a quotient of the disjoint union +It is constructed as a quotient of the free module instead of a quotient of the disjoint union so as to make the operations (addition etc.) "computable". ## Main definitions * `Module.DirectLimit G f` * `AddCommGroup.DirectLimit G f` -* `Ring.DirectLimit G f` -/ @@ -428,326 +424,3 @@ end functorial end DirectLimit end AddCommGroup - -namespace Ring - -variable (G) [∀ i, CommRing (G i)] - -section - -variable (f : ∀ i j, i ≤ j → G i → G j) - -open FreeCommRing - -/-- The direct limit of a directed system is the rings glued together along the maps. -/ -def DirectLimit : Type _ := - FreeCommRing (Σ i, G i) ⧸ - Ideal.span - { a | - (∃ i j H x, of (⟨j, f i j H x⟩ : Σ i, G i) - of ⟨i, x⟩ = a) ∨ - (∃ i, of (⟨i, 1⟩ : Σ i, G i) - 1 = a) ∨ - (∃ i x y, of (⟨i, x + y⟩ : Σ i, G i) - (of ⟨i, x⟩ + of ⟨i, y⟩) = a) ∨ - ∃ i x y, of (⟨i, x * y⟩ : Σ i, G i) - of ⟨i, x⟩ * of ⟨i, y⟩ = a } - -namespace DirectLimit - -instance commRing : CommRing (DirectLimit G f) := - Ideal.Quotient.commRing _ - -instance ring : Ring (DirectLimit G f) := - CommRing.toRing - --- Porting note: Added a `Zero` instance to get rid of `0` errors. -instance zero : Zero (DirectLimit G f) := by - unfold DirectLimit - exact ⟨0⟩ - -instance : Inhabited (DirectLimit G f) := - ⟨0⟩ - -/-- The canonical map from a component to the direct limit. -/ -nonrec def of (i) : G i →+* DirectLimit G f := - RingHom.mk' - { toFun := fun x ↦ Ideal.Quotient.mk _ (of (⟨i, x⟩ : Σ i, G i)) - map_one' := Ideal.Quotient.eq.2 <| subset_span <| Or.inr <| Or.inl ⟨i, rfl⟩ - map_mul' := fun x y ↦ - Ideal.Quotient.eq.2 <| subset_span <| Or.inr <| Or.inr <| Or.inr ⟨i, x, y, rfl⟩ } - fun x y ↦ Ideal.Quotient.eq.2 <| subset_span <| Or.inr <| Or.inr <| Or.inl ⟨i, x, y, rfl⟩ - -variable {G f} - -theorem quotientMk_of (i x) : Ideal.Quotient.mk _ (.of ⟨i, x⟩) = of G f i x := - rfl - -@[simp] theorem of_f {i j} (hij) (x) : of G f j (f i j hij x) = of G f i x := - Ideal.Quotient.eq.2 <| subset_span <| Or.inl ⟨i, j, hij, x, rfl⟩ - -/-- Every element of the direct limit corresponds to some element in -some component of the directed system. -/ -theorem exists_of [Nonempty ι] [IsDirected ι (· ≤ ·)] (z : DirectLimit G f) : - ∃ i x, of G f i x = z := by - obtain ⟨z, rfl⟩ := Ideal.Quotient.mk_surjective z - refine z.induction_on ⟨Classical.arbitrary ι, -1, by simp⟩ (fun ⟨i, x⟩ ↦ ⟨i, x, rfl⟩) ?_ ?_ <;> - rintro x' y' ⟨i, x, hx⟩ ⟨j, y, hy⟩ <;> have ⟨k, hik, hjk⟩ := exists_ge_ge i j - · exact ⟨k, f i k hik x + f j k hjk y, by rw [map_add, of_f, of_f, hx, hy]; rfl⟩ - · exact ⟨k, f i k hik x * f j k hjk y, by rw [map_mul, of_f, of_f, hx, hy]; rfl⟩ - -section - -open Polynomial - -variable {f' : ∀ i j, i ≤ j → G i →+* G j} - -nonrec theorem Polynomial.exists_of [Nonempty ι] [IsDirected ι (· ≤ ·)] - (q : Polynomial (DirectLimit G fun i j h ↦ f' i j h)) : - ∃ i p, Polynomial.map (of G (fun i j h ↦ f' i j h) i) p = q := - Polynomial.induction_on q - (fun z ↦ - let ⟨i, x, h⟩ := exists_of z - ⟨i, C x, by rw [map_C, h]⟩) - (fun q₁ q₂ ⟨i₁, p₁, ih₁⟩ ⟨i₂, p₂, ih₂⟩ ↦ - let ⟨i, h1, h2⟩ := exists_ge_ge i₁ i₂ - ⟨i, p₁.map (f' i₁ i h1) + p₂.map (f' i₂ i h2), by - rw [Polynomial.map_add, map_map, map_map, ← ih₁, ← ih₂] - congr 2 <;> ext x <;> simp_rw [RingHom.comp_apply, of_f]⟩) - fun n z _ ↦ - let ⟨i, x, h⟩ := exists_of z - ⟨i, C x * X ^ (n + 1), by rw [Polynomial.map_mul, map_C, h, Polynomial.map_pow, map_X]⟩ - -end - -@[elab_as_elim] -theorem induction_on [Nonempty ι] [IsDirected ι (· ≤ ·)] {C : DirectLimit G f → Prop} - (z : DirectLimit G f) (ih : ∀ i x, C (of G f i x)) : C z := - let ⟨i, x, hx⟩ := exists_of z - hx ▸ ih i x - -variable (P : Type*) [CommRing P] - -open FreeCommRing - -variable (G f) in -/-- The universal property of the direct limit: maps from the components to another ring -that respect the directed system structure (i.e. make some diagram commute) give rise -to a unique map out of the direct limit. --/ -def lift (g : ∀ i, G i →+* P) (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) : - DirectLimit G f →+* P := - Ideal.Quotient.lift _ (FreeCommRing.lift fun x : Σ i, G i ↦ g x.1 x.2) - (by - suffices Ideal.span _ ≤ - Ideal.comap (FreeCommRing.lift fun x : Σ i : ι, G i ↦ g x.fst x.snd) ⊥ by - intro x hx - exact (mem_bot P).1 (this hx) - rw [Ideal.span_le] - intro x hx - rw [SetLike.mem_coe, Ideal.mem_comap, mem_bot] - rcases hx with (⟨i, j, hij, x, rfl⟩ | ⟨i, rfl⟩ | ⟨i, x, y, rfl⟩ | ⟨i, x, y, rfl⟩) <;> - simp only [RingHom.map_sub, lift_of, Hg, RingHom.map_one, RingHom.map_add, RingHom.map_mul, - (g i).map_one, (g i).map_add, (g i).map_mul, sub_self]) - -variable (g : ∀ i, G i →+* P) (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) - -@[simp] theorem lift_of (i x) : lift G f P g Hg (of G f i x) = g i x := - FreeCommRing.lift_of _ _ - -theorem lift_unique (F : DirectLimit G f →+* P) (x) : - F x = lift G f P (fun i ↦ F.comp <| of G f i) (fun i j hij x ↦ by simp) x := by - obtain ⟨x, rfl⟩ := Ideal.Quotient.mk_surjective x - exact x.induction_on (by simp) (fun _ ↦ .symm <| lift_of ..) - (by simp +contextual) (by simp+contextual) - -lemma lift_injective [Nonempty ι] [IsDirected ι (· ≤ ·)] - (injective : ∀ i, Function.Injective <| g i) : - Function.Injective (lift G f P g Hg) := by - simp_rw [injective_iff_map_eq_zero] at injective ⊢ - intros z hz - induction z using DirectLimit.induction_on with - | ih _ g => rw [lift_of] at hz; rw [injective _ g hz, _root_.map_zero] - -section OfZeroExact - -variable (f' : ∀ i j, i ≤ j → G i →+* G j) -variable [DirectedSystem G fun i j h ↦ f' i j h] [IsDirected ι (· ≤ ·)] -variable (G f) - -open _root_.DirectLimit in -/-- The direct limit constructed as a quotient of the free commutative ring is isomorphic to -the direct limit constructed as a quotient of the disjoint union. -/ -def ringEquiv [Nonempty ι] : DirectLimit G (f' · · ·) ≃+* _root_.DirectLimit G f' := - .ofRingHom (lift _ _ _ (Ring.of _ _) fun _ _ _ _ ↦ .symm <| eq_of_le ..) - (Ring.lift _ _ _ (of _ _) fun _ _ _ _ ↦ of_f ..) - (by ext ⟨_⟩; rw [← Quotient.mk]; simp [Ring.lift, _root_.DirectLimit.lift_def]; rfl) - (by ext x; exact x.induction_on fun i x ↦ by simp) - -theorem ringEquiv_of [Nonempty ι] {i g} : ringEquiv G f' (of _ _ i g) = ⟦⟨i, g⟩⟧ := by - simp [ringEquiv]; rfl - -theorem ringEquiv_symm_mk [Nonempty ι] {g} : (ringEquiv G f').symm ⟦g⟧ = of _ _ g.1 g.2 := rfl - -variable {G f'} -/-- A component that corresponds to zero in the direct limit is already zero in some -bigger module in the directed system. -/ -theorem of.zero_exact {i x} (hix : of G (f' · · ·) i x = 0) : - ∃ (j : _) (hij : i ≤ j), f' i j hij x = 0 := by - haveI := Nonempty.intro i - apply_fun ringEquiv _ _ at hix - rwa [map_zero, ringEquiv_of, DirectLimit.exists_eq_zero] at hix - -end OfZeroExact - -variable (f' : ∀ i j, i ≤ j → G i →+* G j) - -/-- If the maps in the directed system are injective, then the canonical maps -from the components to the direct limits are injective. -/ -theorem of_injective [IsDirected ι (· ≤ ·)] [DirectedSystem G fun i j h ↦ f' i j h] - (hf : ∀ i j hij, Function.Injective (f' i j hij)) (i) : - Function.Injective (of G (fun i j h ↦ f' i j h) i) := - haveI := Nonempty.intro i - ((ringEquiv _ _).comp_injective _).mp - fun _ _ eq ↦ DirectLimit.mk_injective f' hf _ (by simpa only [← ringEquiv_of]) - -section functorial - -variable {f : ∀ i j, i ≤ j → G i →+* G j} -variable {G' : ι → Type*} [∀ i, CommRing (G' i)] -variable {f' : ∀ i j, i ≤ j → G' i →+* G' j} -variable {G'' : ι → Type*} [∀ i, CommRing (G'' i)] -variable {f'' : ∀ i j, i ≤ j → G'' i →+* G'' j} - -/-- -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 -homomorphism `lim G ⟶ lim G'`. --/ -def map (g : (i : ι) → G i →+* G' i) - (hg : ∀ i j h, (g j).comp (f i j h) = (f' i j h).comp (g i)) : - DirectLimit G (fun _ _ h ↦ f _ _ h) →+* DirectLimit G' fun _ _ h ↦ f' _ _ h := - lift _ _ _ (fun i ↦ (of _ _ _).comp (g i)) fun i j h g ↦ by - have eq1 := DFunLike.congr_fun (hg i j h) g - simp only [RingHom.coe_comp, Function.comp_apply] at eq1 ⊢ - rw [eq1, of_f] - -@[simp] lemma map_apply_of (g : (i : ι) → G i →+* G' i) - (hg : ∀ i j h, (g j).comp (f i j h) = (f' i j h).comp (g i)) - {i : ι} (x : G i) : - map g hg (of G _ _ x) = of G' (fun _ _ h ↦ f' _ _ h) i (g i x) := - lift_of _ _ _ _ _ - -@[simp] lemma map_id : - map (fun _ ↦ RingHom.id _) (fun _ _ _ ↦ rfl) = RingHom.id (DirectLimit G fun _ _ h ↦ f _ _ h) := - DFunLike.ext _ _ fun x ↦ by - obtain ⟨x, rfl⟩ := Ideal.Quotient.mk_surjective x - refine x.induction_on (by simp) (fun _ ↦ ?_) (by simp+contextual) (by simp+contextual) - rw [quotientMk_of, map_apply_of]; rfl - -lemma map_comp (g₁ : (i : ι) → G i →+* G' i) (g₂ : (i : ι) → G' i →+* G'' i) - (hg₁ : ∀ i j h, (g₁ j).comp (f i j h) = (f' i j h).comp (g₁ i)) - (hg₂ : ∀ i j h, (g₂ j).comp (f' i j h) = (f'' i j h).comp (g₂ i)) : - ((map g₂ hg₂).comp (map g₁ hg₁) : - DirectLimit G (fun _ _ h ↦ f _ _ h) →+* DirectLimit G'' fun _ _ h ↦ f'' _ _ h) = - (map (fun i ↦ (g₂ i).comp (g₁ i)) fun i j h ↦ by - rw [RingHom.comp_assoc, hg₁ i, ← RingHom.comp_assoc, hg₂ i, RingHom.comp_assoc] : - DirectLimit G (fun _ _ h ↦ f _ _ h) →+* DirectLimit G'' fun _ _ h ↦ f'' _ _ h) := - DFunLike.ext _ _ fun x ↦ by - obtain ⟨x, rfl⟩ := Ideal.Quotient.mk_surjective x - refine x.induction_on (by simp) (fun _ ↦ ?_) (by simp+contextual) (by simp+contextual) - rw [RingHom.comp_apply, quotientMk_of] - simp_rw [map_apply_of] - rfl - -/-- -Consider direct limits `lim G` and `lim G'` with direct system `f` and `f'` respectively, any -family of equivalences `eᵢ : Gᵢ ≅ G'ᵢ` such that `e ∘ f = f' ∘ e` induces an equivalence -`lim G ⟶ lim G'`. --/ -def congr (e : (i : ι) → G i ≃+* G' i) - (he : ∀ i j h, (e j).toRingHom.comp (f i j h) = (f' i j h).comp (e i)) : - DirectLimit G (fun _ _ h ↦ f _ _ h) ≃+* DirectLimit G' fun _ _ h ↦ f' _ _ h := - RingEquiv.ofRingHom - (map (e ·) he) - (map (fun i ↦ (e i).symm) fun i j h ↦ DFunLike.ext _ _ fun x ↦ by - have eq1 := DFunLike.congr_fun (he i j h) ((e i).symm x) - simp only [RingEquiv.toRingHom_eq_coe, RingHom.coe_comp, RingHom.coe_coe, Function.comp_apply, - RingEquiv.apply_symm_apply] at eq1 ⊢ - simp [← eq1, of_f]) - (by simp [map_comp]) (by simp [map_comp]) - -lemma congr_apply_of (e : (i : ι) → G i ≃+* G' i) - (he : ∀ i j h, (e j).toRingHom.comp (f i j h) = (f' i j h).comp (e i)) - {i : ι} (g : G i) : - congr e he (of G _ i g) = of G' (fun _ _ h ↦ f' _ _ h) i (e i g) := - map_apply_of _ he _ - -lemma congr_symm_apply_of (e : (i : ι) → G i ≃+* G' i) - (he : ∀ i j h, (e j).toRingHom.comp (f i j h) = (f' i j h).comp (e i)) - {i : ι} (g : G' i) : - (congr e he).symm (of G' _ i g) = of G (fun _ _ h ↦ f _ _ h) i ((e i).symm g) := by - simp only [congr, RingEquiv.ofRingHom_symm_apply, map_apply_of, RingHom.coe_coe] - -end functorial - -end DirectLimit - -end - -end Ring - -namespace Field - -variable [Nonempty ι] [IsDirected ι (· ≤ ·)] [∀ i, Field (G i)] (G) -variable (f : ∀ i j, i ≤ j → G i → G j) -variable (f' : ∀ i j, i ≤ j → G i →+* G j) - -namespace DirectLimit - -instance nontrivial [DirectedSystem G (f' · · ·)] : - Nontrivial (Ring.DirectLimit G (f' · · ·)) := - ⟨⟨0, 1, - Nonempty.elim (by infer_instance) fun i : ι ↦ by - change (0 : Ring.DirectLimit G (f' · · ·)) ≠ 1 - rw [← (Ring.DirectLimit.of _ _ _).map_one] - · intro H; rcases Ring.DirectLimit.of.zero_exact H.symm with ⟨j, hij, hf⟩ - rw [(f' i j hij).map_one] at hf - exact one_ne_zero hf⟩⟩ - -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 - rw [← (Ring.DirectLimit.of _ _ _).map_mul, - mul_inv_cancel₀ fun h : x = 0 ↦ H <| by rw [h, (Ring.DirectLimit.of _ _ _).map_zero], - (Ring.DirectLimit.of _ _ _).map_one]⟩ - -section - - -open Classical in -/-- Noncomputable multiplicative inverse in a direct limit of fields. -/ -noncomputable def inv (p : Ring.DirectLimit G f) : Ring.DirectLimit G f := - if H : p = 0 then 0 else Classical.choose (DirectLimit.exists_inv G f H) - -protected theorem mul_inv_cancel {p : Ring.DirectLimit G f} (hp : p ≠ 0) : p * inv G f p = 1 := by - rw [inv, dif_neg hp, Classical.choose_spec (DirectLimit.exists_inv G f hp)] - -protected theorem inv_mul_cancel {p : Ring.DirectLimit G f} (hp : p ≠ 0) : inv G f p * p = 1 := by - rw [_root_.mul_comm, DirectLimit.mul_inv_cancel G f hp] - -/-- Noncomputable field structure on the direct limit of fields. -See note [reducible non-instances]. -/ -protected noncomputable abbrev field [DirectedSystem G (f' · · ·)] : - Field (Ring.DirectLimit G (f' · · ·)) where - -- This used to include the parent CommRing and Nontrivial instances, - -- but leaving them implicit avoids a very expensive (2-3 minutes!) eta expansion. - inv := inv G (f' · · ·) - mul_inv_cancel := fun _ ↦ DirectLimit.mul_inv_cancel G (f' · · ·) - inv_zero := dif_pos rfl - nnqsmul := _ - nnqsmul_def _ _ := rfl - qsmul := _ - qsmul_def _ _ := rfl - -end - -end DirectLimit - -end Field diff --git a/Mathlib/Algebra/Colimit/Ring.lean b/Mathlib/Algebra/Colimit/Ring.lean new file mode 100644 index 0000000000000..0ee9c31b5230d --- /dev/null +++ b/Mathlib/Algebra/Colimit/Ring.lean @@ -0,0 +1,356 @@ +/- +Copyright (c) 2019 Kenny Lau, Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kenny Lau, Chris Hughes, Jujian Zhang +-/ +import Mathlib.Data.Finset.Order +import Mathlib.RingTheory.FreeCommRing +import Mathlib.RingTheory.Ideal.Maps +import Mathlib.RingTheory.Ideal.Quotient.Defs +import Mathlib.Algebra.Colimit.DirectLimit +import Mathlib.Tactic.SuppressCompilation + +/-! +# Direct limit of rings, and fields + +See Atiyah-Macdonald PP.32-33, Matsumura PP.269-270 + +Generalizes the notion of "union", or "gluing", of incomparable rings or fields. + +It is constructed as a quotient of the free commutative ring instead of a quotient of +the disjoint union so as to make the operations (addition etc.) "computable". + +## Main definition + +* `Ring.DirectLimit G f` + +-/ + +suppress_compilation + +variable {ι : Type*} [Preorder ι] (G : ι → Type*) + +open Submodule + +namespace Ring + +variable [∀ i, CommRing (G i)] + +section + +variable (f : ∀ i j, i ≤ j → G i → G j) + +open FreeCommRing + +/-- The direct limit of a directed system is the rings glued together along the maps. -/ +def DirectLimit : Type _ := + FreeCommRing (Σ i, G i) ⧸ + Ideal.span + { a | + (∃ i j H x, of (⟨j, f i j H x⟩ : Σ i, G i) - of ⟨i, x⟩ = a) ∨ + (∃ i, of (⟨i, 1⟩ : Σ i, G i) - 1 = a) ∨ + (∃ i x y, of (⟨i, x + y⟩ : Σ i, G i) - (of ⟨i, x⟩ + of ⟨i, y⟩) = a) ∨ + ∃ i x y, of (⟨i, x * y⟩ : Σ i, G i) - of ⟨i, x⟩ * of ⟨i, y⟩ = a } + +namespace DirectLimit + +instance commRing : CommRing (DirectLimit G f) := + Ideal.Quotient.commRing _ + +instance ring : Ring (DirectLimit G f) := + CommRing.toRing + +-- Porting note: Added a `Zero` instance to get rid of `0` errors. +instance zero : Zero (DirectLimit G f) := by + unfold DirectLimit + exact ⟨0⟩ + +instance : Inhabited (DirectLimit G f) := + ⟨0⟩ + +/-- The canonical map from a component to the direct limit. -/ +nonrec def of (i) : G i →+* DirectLimit G f := + RingHom.mk' + { toFun := fun x ↦ Ideal.Quotient.mk _ (of (⟨i, x⟩ : Σ i, G i)) + map_one' := Ideal.Quotient.eq.2 <| subset_span <| Or.inr <| Or.inl ⟨i, rfl⟩ + map_mul' := fun x y ↦ + Ideal.Quotient.eq.2 <| subset_span <| Or.inr <| Or.inr <| Or.inr ⟨i, x, y, rfl⟩ } + fun x y ↦ Ideal.Quotient.eq.2 <| subset_span <| Or.inr <| Or.inr <| Or.inl ⟨i, x, y, rfl⟩ + +variable {G f} + +theorem quotientMk_of (i x) : Ideal.Quotient.mk _ (.of ⟨i, x⟩) = of G f i x := + rfl + +@[simp] theorem of_f {i j} (hij) (x) : of G f j (f i j hij x) = of G f i x := + Ideal.Quotient.eq.2 <| subset_span <| Or.inl ⟨i, j, hij, x, rfl⟩ + +/-- Every element of the direct limit corresponds to some element in +some component of the directed system. -/ +theorem exists_of [Nonempty ι] [IsDirected ι (· ≤ ·)] (z : DirectLimit G f) : + ∃ i x, of G f i x = z := by + obtain ⟨z, rfl⟩ := Ideal.Quotient.mk_surjective z + refine z.induction_on ⟨Classical.arbitrary ι, -1, by simp⟩ (fun ⟨i, x⟩ ↦ ⟨i, x, rfl⟩) ?_ ?_ <;> + rintro x' y' ⟨i, x, hx⟩ ⟨j, y, hy⟩ <;> have ⟨k, hik, hjk⟩ := exists_ge_ge i j + · exact ⟨k, f i k hik x + f j k hjk y, by rw [map_add, of_f, of_f, hx, hy]; rfl⟩ + · exact ⟨k, f i k hik x * f j k hjk y, by rw [map_mul, of_f, of_f, hx, hy]; rfl⟩ + +section + +open Polynomial + +variable {f' : ∀ i j, i ≤ j → G i →+* G j} + +nonrec theorem Polynomial.exists_of [Nonempty ι] [IsDirected ι (· ≤ ·)] + (q : Polynomial (DirectLimit G fun i j h ↦ f' i j h)) : + ∃ i p, Polynomial.map (of G (fun i j h ↦ f' i j h) i) p = q := + Polynomial.induction_on q + (fun z ↦ + let ⟨i, x, h⟩ := exists_of z + ⟨i, C x, by rw [map_C, h]⟩) + (fun q₁ q₂ ⟨i₁, p₁, ih₁⟩ ⟨i₂, p₂, ih₂⟩ ↦ + let ⟨i, h1, h2⟩ := exists_ge_ge i₁ i₂ + ⟨i, p₁.map (f' i₁ i h1) + p₂.map (f' i₂ i h2), by + rw [Polynomial.map_add, map_map, map_map, ← ih₁, ← ih₂] + congr 2 <;> ext x <;> simp_rw [RingHom.comp_apply, of_f]⟩) + fun n z _ ↦ + let ⟨i, x, h⟩ := exists_of z + ⟨i, C x * X ^ (n + 1), by rw [Polynomial.map_mul, map_C, h, Polynomial.map_pow, map_X]⟩ + +end + +@[elab_as_elim] +theorem induction_on [Nonempty ι] [IsDirected ι (· ≤ ·)] {C : DirectLimit G f → Prop} + (z : DirectLimit G f) (ih : ∀ i x, C (of G f i x)) : C z := + let ⟨i, x, hx⟩ := exists_of z + hx ▸ ih i x + +variable (P : Type*) [CommRing P] + +open FreeCommRing + +variable (G f) in +/-- The universal property of the direct limit: maps from the components to another ring +that respect the directed system structure (i.e. make some diagram commute) give rise +to a unique map out of the direct limit. +-/ +def lift (g : ∀ i, G i →+* P) (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) : + DirectLimit G f →+* P := + Ideal.Quotient.lift _ (FreeCommRing.lift fun x : Σ i, G i ↦ g x.1 x.2) + (by + suffices Ideal.span _ ≤ + Ideal.comap (FreeCommRing.lift fun x : Σ i : ι, G i ↦ g x.fst x.snd) ⊥ by + intro x hx + exact (mem_bot P).1 (this hx) + rw [Ideal.span_le] + intro x hx + rw [SetLike.mem_coe, Ideal.mem_comap, mem_bot] + rcases hx with (⟨i, j, hij, x, rfl⟩ | ⟨i, rfl⟩ | ⟨i, x, y, rfl⟩ | ⟨i, x, y, rfl⟩) <;> + simp only [RingHom.map_sub, lift_of, Hg, RingHom.map_one, RingHom.map_add, RingHom.map_mul, + (g i).map_one, (g i).map_add, (g i).map_mul, sub_self]) + +variable (g : ∀ i, G i →+* P) (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) + +@[simp] theorem lift_of (i x) : lift G f P g Hg (of G f i x) = g i x := + FreeCommRing.lift_of _ _ + +theorem lift_unique (F : DirectLimit G f →+* P) (x) : + F x = lift G f P (fun i ↦ F.comp <| of G f i) (fun i j hij x ↦ by simp) x := by + obtain ⟨x, rfl⟩ := Ideal.Quotient.mk_surjective x + exact x.induction_on (by simp) (fun _ ↦ .symm <| lift_of ..) + (by simp +contextual) (by simp+contextual) + +lemma lift_injective [Nonempty ι] [IsDirected ι (· ≤ ·)] + (injective : ∀ i, Function.Injective <| g i) : + Function.Injective (lift G f P g Hg) := by + simp_rw [injective_iff_map_eq_zero] at injective ⊢ + intros z hz + induction z using DirectLimit.induction_on with + | ih _ g => rw [lift_of] at hz; rw [injective _ g hz, _root_.map_zero] + +section OfZeroExact + +variable (f' : ∀ i j, i ≤ j → G i →+* G j) +variable [DirectedSystem G fun i j h ↦ f' i j h] [IsDirected ι (· ≤ ·)] +variable (G f) + +open _root_.DirectLimit in +/-- The direct limit constructed as a quotient of the free commutative ring is isomorphic to +the direct limit constructed as a quotient of the disjoint union. -/ +def ringEquiv [Nonempty ι] : DirectLimit G (f' · · ·) ≃+* _root_.DirectLimit G f' := + .ofRingHom (lift _ _ _ (Ring.of _ _) fun _ _ _ _ ↦ .symm <| eq_of_le ..) + (Ring.lift _ _ _ (of _ _) fun _ _ _ _ ↦ of_f ..) + (by ext ⟨_⟩; rw [← Quotient.mk]; simp [Ring.lift, _root_.DirectLimit.lift_def]; rfl) + (by ext x; exact x.induction_on fun i x ↦ by simp) + +theorem ringEquiv_of [Nonempty ι] {i g} : ringEquiv G f' (of _ _ i g) = ⟦⟨i, g⟩⟧ := by + simp [ringEquiv]; rfl + +theorem ringEquiv_symm_mk [Nonempty ι] {g} : (ringEquiv G f').symm ⟦g⟧ = of _ _ g.1 g.2 := rfl + +variable {G f'} +/-- A component that corresponds to zero in the direct limit is already zero in some +bigger module in the directed system. -/ +theorem of.zero_exact {i x} (hix : of G (f' · · ·) i x = 0) : + ∃ (j : _) (hij : i ≤ j), f' i j hij x = 0 := by + haveI := Nonempty.intro i + apply_fun ringEquiv _ _ at hix + rwa [map_zero, ringEquiv_of, DirectLimit.exists_eq_zero] at hix + +end OfZeroExact + +variable (f' : ∀ i j, i ≤ j → G i →+* G j) + +/-- If the maps in the directed system are injective, then the canonical maps +from the components to the direct limits are injective. -/ +theorem of_injective [IsDirected ι (· ≤ ·)] [DirectedSystem G fun i j h ↦ f' i j h] + (hf : ∀ i j hij, Function.Injective (f' i j hij)) (i) : + Function.Injective (of G (fun i j h ↦ f' i j h) i) := + haveI := Nonempty.intro i + ((ringEquiv _ _).comp_injective _).mp + fun _ _ eq ↦ DirectLimit.mk_injective f' hf _ (by simpa only [← ringEquiv_of]) + +section functorial + +variable {f : ∀ i j, i ≤ j → G i →+* G j} +variable {G' : ι → Type*} [∀ i, CommRing (G' i)] +variable {f' : ∀ i j, i ≤ j → G' i →+* G' j} +variable {G'' : ι → Type*} [∀ i, CommRing (G'' i)] +variable {f'' : ∀ i j, i ≤ j → G'' i →+* G'' j} + +/-- +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 +homomorphism `lim G ⟶ lim G'`. +-/ +def map (g : (i : ι) → G i →+* G' i) + (hg : ∀ i j h, (g j).comp (f i j h) = (f' i j h).comp (g i)) : + DirectLimit G (fun _ _ h ↦ f _ _ h) →+* DirectLimit G' fun _ _ h ↦ f' _ _ h := + lift _ _ _ (fun i ↦ (of _ _ _).comp (g i)) fun i j h g ↦ by + have eq1 := DFunLike.congr_fun (hg i j h) g + simp only [RingHom.coe_comp, Function.comp_apply] at eq1 ⊢ + rw [eq1, of_f] + +@[simp] lemma map_apply_of (g : (i : ι) → G i →+* G' i) + (hg : ∀ i j h, (g j).comp (f i j h) = (f' i j h).comp (g i)) + {i : ι} (x : G i) : + map g hg (of G _ _ x) = of G' (fun _ _ h ↦ f' _ _ h) i (g i x) := + lift_of _ _ _ _ _ + +@[simp] lemma map_id : + map (fun _ ↦ RingHom.id _) (fun _ _ _ ↦ rfl) = RingHom.id (DirectLimit G fun _ _ h ↦ f _ _ h) := + DFunLike.ext _ _ fun x ↦ by + obtain ⟨x, rfl⟩ := Ideal.Quotient.mk_surjective x + refine x.induction_on (by simp) (fun _ ↦ ?_) (by simp+contextual) (by simp+contextual) + rw [quotientMk_of, map_apply_of]; rfl + +lemma map_comp (g₁ : (i : ι) → G i →+* G' i) (g₂ : (i : ι) → G' i →+* G'' i) + (hg₁ : ∀ i j h, (g₁ j).comp (f i j h) = (f' i j h).comp (g₁ i)) + (hg₂ : ∀ i j h, (g₂ j).comp (f' i j h) = (f'' i j h).comp (g₂ i)) : + ((map g₂ hg₂).comp (map g₁ hg₁) : + DirectLimit G (fun _ _ h ↦ f _ _ h) →+* DirectLimit G'' fun _ _ h ↦ f'' _ _ h) = + (map (fun i ↦ (g₂ i).comp (g₁ i)) fun i j h ↦ by + rw [RingHom.comp_assoc, hg₁ i, ← RingHom.comp_assoc, hg₂ i, RingHom.comp_assoc] : + DirectLimit G (fun _ _ h ↦ f _ _ h) →+* DirectLimit G'' fun _ _ h ↦ f'' _ _ h) := + DFunLike.ext _ _ fun x ↦ by + obtain ⟨x, rfl⟩ := Ideal.Quotient.mk_surjective x + refine x.induction_on (by simp) (fun _ ↦ ?_) (by simp+contextual) (by simp+contextual) + rw [RingHom.comp_apply, quotientMk_of] + simp_rw [map_apply_of] + rfl + +/-- +Consider direct limits `lim G` and `lim G'` with direct system `f` and `f'` respectively, any +family of equivalences `eᵢ : Gᵢ ≅ G'ᵢ` such that `e ∘ f = f' ∘ e` induces an equivalence +`lim G ⟶ lim G'`. +-/ +def congr (e : (i : ι) → G i ≃+* G' i) + (he : ∀ i j h, (e j).toRingHom.comp (f i j h) = (f' i j h).comp (e i)) : + DirectLimit G (fun _ _ h ↦ f _ _ h) ≃+* DirectLimit G' fun _ _ h ↦ f' _ _ h := + RingEquiv.ofRingHom + (map (e ·) he) + (map (fun i ↦ (e i).symm) fun i j h ↦ DFunLike.ext _ _ fun x ↦ by + have eq1 := DFunLike.congr_fun (he i j h) ((e i).symm x) + simp only [RingEquiv.toRingHom_eq_coe, RingHom.coe_comp, RingHom.coe_coe, Function.comp_apply, + RingEquiv.apply_symm_apply] at eq1 ⊢ + simp [← eq1, of_f]) + (by simp [map_comp]) (by simp [map_comp]) + +lemma congr_apply_of (e : (i : ι) → G i ≃+* G' i) + (he : ∀ i j h, (e j).toRingHom.comp (f i j h) = (f' i j h).comp (e i)) + {i : ι} (g : G i) : + congr e he (of G _ i g) = of G' (fun _ _ h ↦ f' _ _ h) i (e i g) := + map_apply_of _ he _ + +lemma congr_symm_apply_of (e : (i : ι) → G i ≃+* G' i) + (he : ∀ i j h, (e j).toRingHom.comp (f i j h) = (f' i j h).comp (e i)) + {i : ι} (g : G' i) : + (congr e he).symm (of G' _ i g) = of G (fun _ _ h ↦ f _ _ h) i ((e i).symm g) := by + simp only [congr, RingEquiv.ofRingHom_symm_apply, map_apply_of, RingHom.coe_coe] + +end functorial + +end DirectLimit + +end + +end Ring + +namespace Field + +variable [Nonempty ι] [IsDirected ι (· ≤ ·)] [∀ i, Field (G i)] +variable (f : ∀ i j, i ≤ j → G i → G j) +variable (f' : ∀ i j, i ≤ j → G i →+* G j) + +namespace DirectLimit + +instance nontrivial [DirectedSystem G (f' · · ·)] : + Nontrivial (Ring.DirectLimit G (f' · · ·)) := + ⟨⟨0, 1, + Nonempty.elim (by infer_instance) fun i : ι ↦ by + change (0 : Ring.DirectLimit G (f' · · ·)) ≠ 1 + rw [← (Ring.DirectLimit.of _ _ _).map_one] + · intro H; rcases Ring.DirectLimit.of.zero_exact H.symm with ⟨j, hij, hf⟩ + rw [(f' i j hij).map_one] at hf + exact one_ne_zero hf⟩⟩ + +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 + rw [← (Ring.DirectLimit.of _ _ _).map_mul, + mul_inv_cancel₀ fun h : x = 0 ↦ H <| by rw [h, (Ring.DirectLimit.of _ _ _).map_zero], + (Ring.DirectLimit.of _ _ _).map_one]⟩ + +section + + +open Classical in +/-- Noncomputable multiplicative inverse in a direct limit of fields. -/ +noncomputable def inv (p : Ring.DirectLimit G f) : Ring.DirectLimit G f := + if H : p = 0 then 0 else Classical.choose (DirectLimit.exists_inv G f H) + +protected theorem mul_inv_cancel {p : Ring.DirectLimit G f} (hp : p ≠ 0) : p * inv G f p = 1 := by + rw [inv, dif_neg hp, Classical.choose_spec (DirectLimit.exists_inv G f hp)] + +protected theorem inv_mul_cancel {p : Ring.DirectLimit G f} (hp : p ≠ 0) : inv G f p * p = 1 := by + rw [_root_.mul_comm, DirectLimit.mul_inv_cancel G f hp] + +/-- Noncomputable field structure on the direct limit of fields. +See note [reducible non-instances]. -/ +protected noncomputable abbrev field [DirectedSystem G (f' · · ·)] : + Field (Ring.DirectLimit G (f' · · ·)) where + -- This used to include the parent CommRing and Nontrivial instances, + -- but leaving them implicit avoids a very expensive (2-3 minutes!) eta expansion. + inv := inv G (f' · · ·) + mul_inv_cancel := fun _ ↦ DirectLimit.mul_inv_cancel G (f' · · ·) + inv_zero := dif_pos rfl + nnqsmul := _ + nnqsmul_def _ _ := rfl + qsmul := _ + qsmul_def _ _ := rfl + +end + +end DirectLimit + +end Field diff --git a/Mathlib/Algebra/Homology/LocalCohomology.lean b/Mathlib/Algebra/Homology/LocalCohomology.lean index c98cac0916bc0..ebc2b550a5d65 100644 --- a/Mathlib/Algebra/Homology/LocalCohomology.lean +++ b/Mathlib/Algebra/Homology/LocalCohomology.lean @@ -9,6 +9,7 @@ import Mathlib.CategoryTheory.Abelian.Ext import Mathlib.CategoryTheory.Limits.Final import Mathlib.RingTheory.Finiteness.Ideal import Mathlib.RingTheory.Ideal.Basic +import Mathlib.RingTheory.Ideal.Quotient.Defs import Mathlib.RingTheory.Noetherian.Defs /-! diff --git a/Mathlib/LinearAlgebra/TensorProduct/DirectLimit.lean b/Mathlib/LinearAlgebra/TensorProduct/DirectLimit.lean index bc41c773a8af6..6af40acde4346 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/DirectLimit.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/DirectLimit.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ -import Mathlib.Algebra.Colimit.ModuleRing +import Mathlib.Algebra.Colimit.Module +import Mathlib.LinearAlgebra.TensorProduct.Basic /-! # Tensor product and direct limits commute with each other. diff --git a/Mathlib/RingTheory/FreeCommRing.lean b/Mathlib/RingTheory/FreeCommRing.lean index ed04aadb7bec1..af33998c3ea6e 100644 --- a/Mathlib/RingTheory/FreeCommRing.lean +++ b/Mathlib/RingTheory/FreeCommRing.lean @@ -94,7 +94,7 @@ theorem of_ne_one (x : α) : of x ≠ 1 := theorem one_ne_of (x : α) : 1 ≠ of x := FreeAbelianGroup.of_injective.ne <| Multiset.zero_ne_singleton _ --- Porting note: added to ease a proof in `Mathlib.Algebra.Colimit.ModuleRing` +-- Porting note: added to ease a proof in `Mathlib.Algebra.Colimit.Ring` lemma of_cons (a : α) (m : Multiset α) : (FreeAbelianGroup.of (Multiplicative.ofAdd (a ::ₘ m))) = @HMul.hMul _ (FreeCommRing α) (FreeCommRing α) _ (of a) (FreeAbelianGroup.of (Multiplicative.ofAdd m)) := by From b145b11c0682f153b5884ef0d487380baef2ccc1 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Thu, 26 Dec 2024 15:27:05 +0000 Subject: [PATCH 005/189] feat: norm positivity extension proves strict positivity (#10276) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Closes #5265. Co-authored-by: Yaël Dillies --- Mathlib/Analysis/InnerProductSpace/Basic.lean | 2 - .../Analysis/InnerProductSpace/Rayleigh.lean | 3 - Mathlib/Analysis/Normed/Group/Basic.lean | 79 ++++++++++++------- .../SimpleGraph/Triangle/Removal.lean | 4 +- .../Euclidean/Angle/Oriented/Basic.lean | 7 +- Mathlib/Tactic/ReduceModChar.lean | 4 +- Mathlib/Util/Qq.lean | 9 ++- MathlibTest/positivity.lean | 11 ++- 8 files changed, 72 insertions(+), 47 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 377de45952c82..cd34649342a00 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -645,8 +645,6 @@ zero. See also `EuclideanGeometry.dist_inversion_inversion` for inversions aroun 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 diff --git a/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean b/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean index ce57eea606617..472b60778ed68 100644 --- a/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean +++ b/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean @@ -57,8 +57,6 @@ theorem rayleigh_smul (x : E) {c : 𝕜} (hc : c ≠ 0) : rayleighQuotient T (c • x) = rayleighQuotient T x := by by_cases hx : x = 0 · simp [hx] - have : ‖c‖ ≠ 0 := by simp [hc] - have : ‖x‖ ≠ 0 := by simp [hx] field_simp [norm_smul, T.reApplyInnerSelf_smul] ring @@ -149,7 +147,6 @@ theorem eq_smul_self_of_isLocalExtrOn_real (hT : IsSelfAdjoint T) {x₀ : F} apply smul_right_injective F hb simp [c, eq_neg_of_add_eq_zero_left h₂, ← mul_smul, this] convert hc - have : ‖x₀‖ ≠ 0 := by simp [hx₀] have := congr_arg (fun x => ⟪x, x₀⟫_ℝ) hc field_simp [inner_smul_left, real_inner_self_eq_norm_mul_norm, sq] at this ⊢ exact this diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index 8390a64944aeb..e69d345223589 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -428,33 +428,6 @@ attribute [bound] norm_nonneg @[to_additive (attr := simp) abs_norm] theorem abs_norm' (z : E) : |‖z‖| = ‖z‖ := abs_of_nonneg <| norm_nonneg' _ -namespace Mathlib.Meta.Positivity - -open Lean Meta Qq Function - -/-- Extension for the `positivity` tactic: multiplicative norms are nonnegative, via -`norm_nonneg'`. -/ -@[positivity Norm.norm _] -def evalMulNorm : PositivityExt where eval {u α} _zα _pα e := do - match u, α, e with - | 0, ~q(ℝ), ~q(@Norm.norm $β $instDist $a) => - let _inst ← synthInstanceQ q(SeminormedGroup $β) - assertInstancesCommute - pure (.nonnegative q(norm_nonneg' $a)) - | _, _, _ => throwError "not ‖ · ‖" - -/-- Extension for the `positivity` tactic: additive norms are nonnegative, via `norm_nonneg`. -/ -@[positivity Norm.norm _] -def evalAddNorm : PositivityExt where eval {u α} _zα _pα e := do - match u, α, e with - | 0, ~q(ℝ), ~q(@Norm.norm $β $instDist $a) => - let _inst ← synthInstanceQ q(SeminormedAddGroup $β) - assertInstancesCommute - pure (.nonnegative q(norm_nonneg $a)) - | _, _, _ => throwError "not ‖ · ‖" - -end Mathlib.Meta.Positivity - @[to_additive (attr := simp) norm_zero] theorem norm_one' : ‖(1 : E)‖ = 0 := by rw [← dist_one_right, dist_self] @@ -1436,6 +1409,58 @@ alias ⟨_, HasCompactSupport.norm⟩ := hasCompactSupport_norm_iff end NormedAddGroup +/-! ### `positivity` extensions -/ + +namespace Mathlib.Meta.Positivity + +open Lean Meta Qq Function + +/-- Extension for the `positivity` tactic: multiplicative norms are always nonnegative, and positive +on non-one inputs. -/ +@[positivity ‖_‖] +def evalMulNorm : PositivityExt where eval {u α} _ _ e := do + match u, α, e with + | 0, ~q(ℝ), ~q(@Norm.norm $E $_n $a) => + let _seminormedGroup_E ← synthInstanceQ q(SeminormedGroup $E) + assertInstancesCommute + -- Check whether we are in a normed group and whether the context contains a `a ≠ 1` assumption + let o : Option (Q(NormedGroup $E) × Q($a ≠ 1)) := ← do + let .some normedGroup_E ← trySynthInstanceQ q(NormedGroup $E) | return none + let some pa ← findLocalDeclWithTypeQ? q($a ≠ 1) | return none + return some (normedGroup_E, pa) + match o with + -- If so, return a proof of `0 < ‖a‖` + | some (_normedGroup_E, pa) => + assertInstancesCommute + return .positive q(norm_pos_iff'.2 $pa) + -- Else, return a proof of `0 ≤ ‖a‖` + | none => return .nonnegative q(norm_nonneg' $a) + | _, _, _ => throwError "not `‖·‖`" + +/-- Extension for the `positivity` tactic: additive norms are always nonnegative, and positive +on non-zero inputs. -/ +@[positivity ‖_‖] +def evalAddNorm : PositivityExt where eval {u α} _ _ e := do + match u, α, e with + | 0, ~q(ℝ), ~q(@Norm.norm $E $_n $a) => + let _seminormedAddGroup_E ← synthInstanceQ q(SeminormedAddGroup $E) + assertInstancesCommute + -- Check whether we are in a normed group and whether the context contains a `a ≠ 0` assumption + let o : Option (Q(NormedAddGroup $E) × Q($a ≠ 0)) := ← do + let .some normedAddGroup_E ← trySynthInstanceQ q(NormedAddGroup $E) | return none + let some pa ← findLocalDeclWithTypeQ? q($a ≠ 0) | return none + return some (normedAddGroup_E, pa) + match o with + -- If so, return a proof of `0 < ‖a‖` + | some (_normedAddGroup_E, pa) => + assertInstancesCommute + return .positive q(norm_pos_iff.2 $pa) + -- Else, return a proof of `0 ≤ ‖a‖` + | none => return .nonnegative q(norm_nonneg $a) + | _, _, _ => throwError "not `‖·‖`" + +end Mathlib.Meta.Positivity + /-! ### Subgroups of normed groups -/ diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Removal.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Removal.lean index b212ba252ae67..86925769d3f15 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Removal.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Removal.lean @@ -165,9 +165,7 @@ Note this looks for `ε ≤ 1` in the context. -/ def evalTriangleRemovalBound : PositivityExt where eval {u α} _zα _pα e := do match u, α, e with | 0, ~q(ℝ), ~q(triangleRemovalBound $ε) => - let some fvarId ← findLocalDeclWithType? q($ε ≤ 1) - | throwError "`ε ≤ 1` is not available in the local context" - let hε₁ : Q($ε ≤ 1) := .fvar fvarId + let some hε₁ ← findLocalDeclWithTypeQ? q($ε ≤ 1) | failure let .positive hε ← core q(inferInstance) q(inferInstance) ε | failure assertInstancesCommute pure (.positive q(triangleRemovalBound_pos $hε $hε₁)) diff --git a/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean b/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean index 3ce4e18dd756a..6572753371e34 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean @@ -442,10 +442,7 @@ theorem oangle_add {x y z : V} (hx : x ≠ 0) (hy : y ≠ 0) (hz : z ≠ 0) : simp_rw [oangle] rw [← Complex.arg_mul_coe_angle, o.kahler_mul y x z] · congr 1 - convert Complex.arg_real_mul _ (_ : 0 < ‖y‖ ^ 2) using 2 - · norm_cast - · have : 0 < ‖y‖ := by simpa using hy - positivity + exact mod_cast Complex.arg_real_mul _ (by positivity : 0 < ‖y‖ ^ 2) · exact o.kahler_ne_zero hx hy · exact o.kahler_ne_zero hy hz @@ -541,8 +538,6 @@ theorem inner_eq_norm_mul_norm_mul_cos_oangle (x y : V) : ⟪x, y⟫ = ‖x‖ * ‖y‖ * Real.Angle.cos (o.oangle x y) := by by_cases hx : x = 0; · simp [hx] by_cases hy : y = 0; · simp [hy] - have : ‖x‖ ≠ 0 := by simpa using hx - 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] -- TODO(https://github.com/leanprover-community/mathlib4/issues/15486): used to be `field_simp`; replaced by `simp only ...` to speed up diff --git a/Mathlib/Tactic/ReduceModChar.lean b/Mathlib/Tactic/ReduceModChar.lean index 18b8a08cf04ce..d386b56ef496e 100644 --- a/Mathlib/Tactic/ReduceModChar.lean +++ b/Mathlib/Tactic/ReduceModChar.lean @@ -192,9 +192,9 @@ match Expr.getAppFnArgs t with let .some instRing ← trySynthInstanceQ q(Ring $t) | return .failure let n ← mkFreshExprMVarQ q(ℕ) - let .some instCharP ← findLocalDeclWithType? q(CharP $t $n) | return .failure + let .some instCharP ← findLocalDeclWithTypeQ? q(CharP $t $n) | return .failure - return .intLike (← instantiateMVarsQ n) instRing (.fvar instCharP) + return .intLike (← instantiateMVarsQ n) instRing instCharP /-- Given an expression `e`, determine whether it is a numeric expression in characteristic `n`, and if so, reduce `e` modulo `n`. diff --git a/Mathlib/Util/Qq.lean b/Mathlib/Util/Qq.lean index 5232470c0d692..baeda29e5657b 100644 --- a/Mathlib/Util/Qq.lean +++ b/Mathlib/Util/Qq.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kim Morrison, Alex J. Best +Authors: Kim Morrison, Alex J. Best, Yaël Dillies -/ import Mathlib.Init import Qq @@ -28,4 +28,11 @@ def inferTypeQ' (e : Expr) : MetaM ((u : Level) × (α : Q(Type $u)) × Q($α)) theorem QuotedDefEq.rfl {u : Level} {α : Q(Sort u)} {a : Q($α)} : @QuotedDefEq u α a a := ⟨⟩ +/-- Return a local declaration whose type is definitionally equal to `sort`. + +This is a Qq version of `Lean.Meta.findLocalDeclWithType?` -/ +def findLocalDeclWithTypeQ? {u : Level} (sort : Q(Sort u)) : MetaM (Option Q($sort)) := do + let some fvarId ← findLocalDeclWithType? q($sort) | return none + return some <| .fvar fvarId + end Qq diff --git a/MathlibTest/positivity.lean b/MathlibTest/positivity.lean index b3f2b9b7af6ba..80b80dab4bb9e 100644 --- a/MathlibTest/positivity.lean +++ b/MathlibTest/positivity.lean @@ -14,7 +14,7 @@ set_option autoImplicit true open Finset Function Nat NNReal ENNReal -variable {ι α β : Type _} +variable {ι α β E : Type*} /- ## Numeric goals -/ @@ -307,8 +307,13 @@ example : 0 ≤ Real.log 1 := by positivity example : 0 ≤ Real.log 0 := by positivity example : 0 ≤ Real.log (-1) := by positivity -example {V : Type _} [NormedCommGroup V] (x : V) : 0 ≤ ‖x‖ := by positivity -example {V : Type _} [NormedAddCommGroup V] (x : V) : 0 ≤ ‖x‖ := by positivity +example [SeminormedGroup E] {a : E} (_ha : a ≠ 1) : 0 ≤ ‖a‖ := by positivity +example [NormedGroup E] {a : E} : 0 ≤ ‖a‖ := by positivity +example [NormedGroup E] {a : E} (ha : a ≠ 1) : 0 < ‖a‖ := by positivity + +example [SeminormedAddGroup E] {a : E} (_ha : a ≠ 0) : 0 ≤ ‖a‖ := by positivity +example [NormedAddGroup E] {a : E} : 0 ≤ ‖a‖ := by positivity +example [NormedAddGroup E] {a : E} (ha : a ≠ 0) : 0 < ‖a‖ := by positivity example [MetricSpace α] (x y : α) : 0 ≤ dist x y := by positivity example [MetricSpace α] {s : Set α} : 0 ≤ Metric.diam s := by positivity From d5e4a9182e9a23c4dd8bec6380816e7fdcc369f5 Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Thu, 26 Dec 2024 18:55:21 +0000 Subject: [PATCH 006/189] refactor(LinearAlgebra/BilinearMap): Improve bilinear map docstrings (#20125) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I don't find the current docstrings at the top of `LinearAlgebra/BilinearMap` helpful for understanding what the definitions do - in fact the current string for `LinearMap.compl₂` and `LinearMap.compr₂` seem to me to be incorrect. This PR gives what seem to me to be more useful / correct descriptions. --- Mathlib/LinearAlgebra/BilinearMap.lean | 18 +++++++++++++----- 1 file changed, 13 insertions(+), 5 deletions(-) diff --git a/Mathlib/LinearAlgebra/BilinearMap.lean b/Mathlib/LinearAlgebra/BilinearMap.lean index 18969882e8214..176ef9be83217 100644 --- a/Mathlib/LinearAlgebra/BilinearMap.lean +++ b/Mathlib/LinearAlgebra/BilinearMap.lean @@ -19,9 +19,16 @@ commuting actions, and `ρ₁₂ : R →+* R₂` and `σ₁₂ : S →+* S₂`. * `LinearMap.mk₂`: a constructor for bilinear maps, taking an unbundled function together with proof witnesses of bilinearity * `LinearMap.flip`: turns a bilinear map `M × N → P` into `N × M → P` -* `LinearMap.lcomp` and `LinearMap.llcomp`: composition of linear maps as a bilinear map -* `LinearMap.compl₂`: composition of a bilinear map `M × N → P` with a linear map `Q → M` -* `LinearMap.compr₂`: composition of a bilinear map `M × N → P` with a linear map `Q → N` +* `LinearMap.lflip`: given a linear map from `M` to `N →ₗ[R] P`, i.e., a bilinear map `M → N → P`, + change the order of variables and get a linear map from `N` to `M →ₗ[R] P`. +* `LinearMap.lcomp`: composition of a given linear map `M → N` with a linear map `N → P` as + a linear map from `Nₗ →ₗ[R] Pₗ` to `M →ₗ[R] Pₗ` +* `LinearMap.llcomp`: composition of linear maps as a bilinear map from `(M →ₗ[R] N) × (N →ₗ[R] P)` + to `M →ₗ[R] P` +* `LinearMap.compl₂`: composition of a linear map `Q → N` and a bilinear map `M → N → P` to + form a bilinear map `M → Q → P`. +* `LinearMap.compr₂`: composition of a linear map `P → Q` and a bilinear map `M → N → P` to form a + bilinear map `M → N → Q`. * `LinearMap.lsmul`: scalar multiplication as a bilinear map `R × M → M` ## Tags @@ -245,7 +252,8 @@ theorem lflip_apply (m : M) (n : N) : lflip f n m = f m n := rfl variable (R Pₗ) -/-- Composing a linear map `M → N` and a linear map `N → P` to form a linear map `M → P`. -/ +/-- Composing a given linear map `M → N` with a linear map `N → P` as a linear map from +`Nₗ →ₗ[R] Pₗ` to `M →ₗ[R] Pₗ`. -/ def lcomp (f : M →ₗ[R] Nₗ) : (Nₗ →ₗ[R] Pₗ) →ₗ[R] M →ₗ[R] Pₗ := flip <| LinearMap.comp (flip id) f @@ -271,7 +279,7 @@ theorem lcompₛₗ_apply (f : M →ₛₗ[σ₁₂] N) (g : N →ₛₗ[σ₂ variable (R M Nₗ Pₗ) -/-- Composing a linear map `M → N` and a linear map `N → P` to form a linear map `M → P`. -/ +/-- Composing linear maps as a bilinear map from `(M →ₗ[R] N) × (N →ₗ[R] P)` to `M →ₗ[R] P` -/ def llcomp : (Nₗ →ₗ[R] Pₗ) →ₗ[R] (M →ₗ[R] Nₗ) →ₗ[R] M →ₗ[R] Pₗ := flip { toFun := lcomp R Pₗ From 5c237a9b8159a7aca14d21ffb8502f6391f74bfd Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 27 Dec 2024 09:22:41 +0000 Subject: [PATCH 007/189] chore(Process/Stopping): merge `measurableSpace_le` and `measurableSpace_le'` (#20237) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Assume `[IsDirected ι (· ≤ ·)]` instead of `SemilatticeSup ι` or `atTop.NeBot` --- Mathlib/Probability/Process/Stopping.lean | 36 ++++++++++------------- 1 file changed, 15 insertions(+), 21 deletions(-) diff --git a/Mathlib/Probability/Process/Stopping.lean b/Mathlib/Probability/Process/Stopping.lean index eb62d512ac6fa..a9b107e3d88e4 100644 --- a/Mathlib/Probability/Process/Stopping.lean +++ b/Mathlib/Probability/Process/Stopping.lean @@ -323,31 +323,25 @@ theorem measurableSpace_le_of_countable [Countable ι] (hτ : IsStoppingTime f · rintro ⟨_, hx, _⟩ exact hx -theorem measurableSpace_le' [IsCountablyGenerated (atTop : Filter ι)] [(atTop : Filter ι).NeBot] +theorem measurableSpace_le [IsCountablyGenerated (atTop : Filter ι)] [IsDirected ι (· ≤ ·)] (hτ : IsStoppingTime f τ) : hτ.measurableSpace ≤ m := by intro s hs - change ∀ i, MeasurableSet[f i] (s ∩ {ω | τ ω ≤ i}) at hs - obtain ⟨seq : ℕ → ι, h_seq_tendsto⟩ := (atTop : Filter ι).exists_seq_tendsto - rw [(_ : s = ⋃ n, s ∩ {ω | τ ω ≤ seq n})] - · exact MeasurableSet.iUnion fun i => f.le (seq i) _ (hs (seq i)) - · ext ω; constructor <;> rw [Set.mem_iUnion] - · intro hx - suffices ∃ i, τ ω ≤ seq i from ⟨this.choose, hx, this.choose_spec⟩ - rw [tendsto_atTop] at h_seq_tendsto - exact (h_seq_tendsto (τ ω)).exists - · rintro ⟨_, hx, _⟩ - exact hx - -theorem measurableSpace_le {ι} [SemilatticeSup ι] {f : Filtration ι m} {τ : Ω → ι} - [IsCountablyGenerated (atTop : Filter ι)] (hτ : IsStoppingTime f τ) : - hτ.measurableSpace ≤ m := by cases isEmpty_or_nonempty ι · haveI : IsEmpty Ω := ⟨fun ω => IsEmpty.false (τ ω)⟩ - intro s _ - suffices hs : s = ∅ by rw [hs]; exact MeasurableSet.empty - haveI : Unique (Set Ω) := Set.uniqueEmpty - rw [Unique.eq_default s, Unique.eq_default ∅] - exact measurableSpace_le' hτ + apply Subsingleton.measurableSet + · change ∀ i, MeasurableSet[f i] (s ∩ {ω | τ ω ≤ i}) at hs + obtain ⟨seq : ℕ → ι, h_seq_tendsto⟩ := (atTop : Filter ι).exists_seq_tendsto + rw [(_ : s = ⋃ n, s ∩ {ω | τ ω ≤ seq n})] + · exact MeasurableSet.iUnion fun i => f.le (seq i) _ (hs (seq i)) + · ext ω; constructor <;> rw [Set.mem_iUnion] + · intro hx + suffices ∃ i, τ ω ≤ seq i from ⟨this.choose, hx, this.choose_spec⟩ + rw [tendsto_atTop] at h_seq_tendsto + exact (h_seq_tendsto (τ ω)).exists + · rintro ⟨_, hx, _⟩ + exact hx + +@[deprecated (since := "2024-12-25")] alias measurableSpace_le' := measurableSpace_le example {f : Filtration ℕ m} {τ : Ω → ℕ} (hτ : IsStoppingTime f τ) : hτ.measurableSpace ≤ m := hτ.measurableSpace_le From 5ed97c4e394ab1b3d2fc513d06216ffacb1a11e0 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 27 Dec 2024 09:22:42 +0000 Subject: [PATCH 008/189] chore(Integral/Asymptotics): drop noop `_root_.Asymptotics` (#20253) They're inside `namespace Asymptotics`. --- Mathlib/MeasureTheory/Integral/Asymptotics.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Integral/Asymptotics.lean b/Mathlib/MeasureTheory/Integral/Asymptotics.lean index a99edd79274bf..4ec76c13473b6 100644 --- a/Mathlib/MeasureTheory/Integral/Asymptotics.lean +++ b/Mathlib/MeasureTheory/Integral/Asymptotics.lean @@ -39,7 +39,7 @@ variable [MeasurableSpace α] [NormedAddCommGroup F] {μ : Measure α} /-- If `f = O[l] g` on measurably generated `l`, `f` is strongly measurable at `l`, and `g` is integrable at `l`, then `f` is integrable at `l`. -/ -theorem _root_.Asymptotics.IsBigO.integrableAtFilter [IsMeasurablyGenerated l] +theorem IsBigO.integrableAtFilter [IsMeasurablyGenerated l] (hf : f =O[l] g) (hfm : StronglyMeasurableAtFilter f l μ) (hg : IntegrableAtFilter g l μ) : IntegrableAtFilter f l μ := by obtain ⟨C, hC⟩ := hf.bound @@ -50,7 +50,7 @@ theorem _root_.Asymptotics.IsBigO.integrableAtFilter [IsMeasurablyGenerated l] exact (hfg x hx).trans (le_abs_self _) /-- Variant of `MeasureTheory.Integrable.mono` taking `f =O[⊤] (g)` instead of `‖f(x)‖ ≤ ‖g(x)‖` -/ -theorem _root_.Asymptotics.IsBigO.integrable (hfm : AEStronglyMeasurable f μ) +theorem IsBigO.integrable (hfm : AEStronglyMeasurable f μ) (hf : f =O[⊤] g) (hg : Integrable g μ) : Integrable f μ := by rewrite [← integrableAtFilter_top] at * exact hf.integrableAtFilter ⟨univ, univ_mem, hfm.restrict⟩ hg From 6ab855d1ff0623256af93fe89f7133276dbb4be1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 27 Dec 2024 11:03:12 +0000 Subject: [PATCH 009/189] chore: rename `MulEquiv.mulEquivOfUnique` to `MulEquiv.ofUnique` (#20243) --- Mathlib/Algebra/GCDMonoid/Basic.lean | 4 ++-- Mathlib/Algebra/Group/Equiv/Basic.lean | 12 +++++++++--- Mathlib/Algebra/Ring/Equiv.lean | 8 +++++--- Mathlib/CategoryTheory/Limits/Shapes/Types.lean | 2 +- Mathlib/Data/FinEnum.lean | 2 +- Mathlib/GroupTheory/Perm/Cycle/Type.lean | 2 +- Mathlib/LinearAlgebra/Dimension/Free.lean | 2 +- Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean | 2 +- Mathlib/Logic/Equiv/Basic.lean | 2 +- Mathlib/Logic/Equiv/Defs.lean | 6 ++++-- Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean | 2 +- Mathlib/Order/Estimator.lean | 2 +- Mathlib/Order/RelIso/Basic.lean | 12 ++++++++---- .../GroupCohomology/Resolution.lean | 2 +- Mathlib/RingTheory/DedekindDomain/Ideal.lean | 3 +-- Mathlib/SetTheory/Cardinal/Basic.lean | 2 +- Mathlib/SetTheory/Game/Nim.lean | 2 +- Mathlib/SetTheory/Game/Ordinal.lean | 2 +- Mathlib/SetTheory/Ordinal/Basic.lean | 2 +- Mathlib/Topology/Homeomorph.lean | 2 +- 20 files changed, 43 insertions(+), 30 deletions(-) diff --git a/Mathlib/Algebra/GCDMonoid/Basic.lean b/Mathlib/Algebra/GCDMonoid/Basic.lean index b6ff248ab4b41..16298274aed73 100644 --- a/Mathlib/Algebra/GCDMonoid/Basic.lean +++ b/Mathlib/Algebra/GCDMonoid/Basic.lean @@ -815,14 +815,14 @@ section UniqueUnit variable [CancelCommMonoidWithZero α] [Subsingleton αˣ] -- see Note [lower instance priority] -instance (priority := 100) normalizationMonoidOfUniqueUnits : NormalizationMonoid α where +instance (priority := 100) NormalizationMonoid.ofUniqueUnits : NormalizationMonoid α where normUnit _ := 1 normUnit_zero := rfl normUnit_mul _ _ := (mul_one 1).symm normUnit_coe_units _ := Subsingleton.elim _ _ instance uniqueNormalizationMonoidOfUniqueUnits : Unique (NormalizationMonoid α) where - default := normalizationMonoidOfUniqueUnits + default := .ofUniqueUnits uniq := fun ⟨u, _, _, _⟩ => by congr; simp [eq_iff_true_of_subsingleton] instance subsingleton_gcdMonoid_of_unique_units : Subsingleton (GCDMonoid α) := diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index eb01af23dc95e..4d516e138b15b 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -446,14 +446,20 @@ section unique /-- The `MulEquiv` between two monoids with a unique element. -/ @[to_additive "The `AddEquiv` between two `AddMonoid`s with a unique element."] -def mulEquivOfUnique {M N} [Unique M] [Unique N] [Mul M] [Mul N] : M ≃* N := - { Equiv.equivOfUnique M N with map_mul' := fun _ _ => Subsingleton.elim _ _ } +def ofUnique {M N} [Unique M] [Unique N] [Mul M] [Mul N] : M ≃* N := + { Equiv.ofUnique M N with map_mul' := fun _ _ => Subsingleton.elim _ _ } + +@[to_additive (attr := deprecated ofUnique (since := "2024-12-25"))] +alias mulEquivOfUnique := ofUnique + +/-- Alias of `AddEquiv.ofEquiv`. -/ +add_decl_doc AddEquiv.addEquivOfUnique /-- There is a unique monoid homomorphism between two monoids with a unique element. -/ @[to_additive "There is a unique additive monoid homomorphism between two additive monoids with a unique element."] instance {M N} [Unique M] [Unique N] [Mul M] [Mul N] : Unique (M ≃* N) where - default := mulEquivOfUnique + default := ofUnique uniq _ := ext fun _ => Subsingleton.elim _ _ end unique diff --git a/Mathlib/Algebra/Ring/Equiv.lean b/Mathlib/Algebra/Ring/Equiv.lean index caec2cc9263f0..f180f913b77df 100644 --- a/Mathlib/Algebra/Ring/Equiv.lean +++ b/Mathlib/Algebra/Ring/Equiv.lean @@ -344,12 +344,14 @@ end trans section unique /-- The `RingEquiv` between two semirings with a unique element. -/ -def ringEquivOfUnique {M N} [Unique M] [Unique N] [Add M] [Mul M] [Add N] [Mul N] : M ≃+* N := - { AddEquiv.addEquivOfUnique, MulEquiv.mulEquivOfUnique with } +def ofUnique {M N} [Unique M] [Unique N] [Add M] [Mul M] [Add N] [Mul N] : M ≃+* N := + { AddEquiv.ofUnique, MulEquiv.ofUnique with } + +@[deprecated (since := "2024-12-26")] alias ringEquivOfUnique := ofUnique instance {M N} [Unique M] [Unique N] [Add M] [Mul M] [Add N] [Mul N] : Unique (M ≃+* N) where - default := ringEquivOfUnique + default := .ofUnique uniq _ := ext fun _ => Subsingleton.elim _ _ end unique diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean index 7778e848fb9c3..f6dfbe28fd6cb 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean @@ -117,7 +117,7 @@ noncomputable instance : Unique (⊤_ (Type u)) := Unique.mk' _ noncomputable def isTerminalEquivUnique (X : Type u) : IsTerminal X ≃ Unique X := equivOfSubsingletonOfSubsingleton (fun h => ((Iso.toEquiv (terminalIsoIsTerminal h).symm).unique)) - (fun _ => IsTerminal.ofIso terminalIsTerminal (Equiv.toIso (Equiv.equivOfUnique _ _))) + (fun _ => IsTerminal.ofIso terminalIsTerminal (Equiv.toIso (Equiv.ofUnique _ _))) /-- A type is terminal if and only if it is isomorphic to `PUnit`. -/ noncomputable def isTerminalEquivIsoPUnit (X : Type u) : IsTerminal X ≃ (X ≅ PUnit) := by diff --git a/Mathlib/Data/FinEnum.lean b/Mathlib/Data/FinEnum.lean index 2258a3d434c08..3d499baf8ef5b 100644 --- a/Mathlib/Data/FinEnum.lean +++ b/Mathlib/Data/FinEnum.lean @@ -227,7 +227,7 @@ aren't two definitionally differing instances around. -/ def ofIsEmpty [IsEmpty α] : FinEnum α := default instance [Unique α] : Unique (FinEnum α) where - default := ⟨1, Equiv.equivOfUnique α (Fin 1)⟩ + default := ⟨1, Equiv.ofUnique α (Fin 1)⟩ uniq e := by show FinEnum.mk e.1 e.2 = _ congr 1 diff --git a/Mathlib/GroupTheory/Perm/Cycle/Type.lean b/Mathlib/GroupTheory/Perm/Cycle/Type.lean index 491829c0b7160..e9a7dd8bc94db 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Type.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Type.lean @@ -437,7 +437,7 @@ def vectorEquiv : List.Vector G n ≃ vectorsProdEqOne G (n + 1) where /-- Given a vector `v` of length `n` whose product is 1, make a vector of length `n - 1`, by deleting the last entry of `v`. -/ def equivVector : ∀ n, vectorsProdEqOne G n ≃ List.Vector G (n - 1) - | 0 => (equivOfUnique (vectorsProdEqOne G 0) (vectorsProdEqOne G 1)).trans (vectorEquiv G 0).symm + | 0 => (ofUnique (vectorsProdEqOne G 0) (vectorsProdEqOne G 1)).trans (vectorEquiv G 0).symm | (n + 1) => (vectorEquiv G n).symm instance [Fintype G] : Fintype (vectorsProdEqOne G n) := diff --git a/Mathlib/LinearAlgebra/Dimension/Free.lean b/Mathlib/LinearAlgebra/Dimension/Free.lean index e168f462ad1fa..6eee759d4983d 100644 --- a/Mathlib/LinearAlgebra/Dimension/Free.lean +++ b/Mathlib/LinearAlgebra/Dimension/Free.lean @@ -207,7 +207,7 @@ noncomputable def basisUnique (ι : Type*) [Unique ι] Basis ι R M := haveI : Module.Finite R M := Module.finite_of_finrank_pos (_root_.zero_lt_one.trans_le h.symm.le) - (finBasisOfFinrankEq R M h).reindex (Equiv.equivOfUnique _ _) + (finBasisOfFinrankEq R M h).reindex (Equiv.ofUnique _ _) @[simp] theorem basisUnique_repr_eq_zero_iff {ι : Type*} [Unique ι] diff --git a/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean b/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean index e6d92885b31c9..1bc8dac9cd047 100644 --- a/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean +++ b/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean @@ -294,7 +294,7 @@ noncomputable def center_equiv_rootsOfUnity : (fun hn ↦ by rw [center_eq_bot_of_subsingleton, Fintype.card_eq_zero, max_eq_right_of_lt zero_lt_one, rootsOfUnity_one] - exact MulEquiv.mulEquivOfUnique) + exact MulEquiv.ofUnique) (fun _ ↦ (max_eq_left (NeZero.one_le : 1 ≤ Fintype.card n)).symm ▸ center_equiv_rootsOfUnity' (Classical.arbitrary n)) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index d09076575b683..093505a183bd2 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1796,7 +1796,7 @@ def uniqueUniqueEquiv : Unique (Unique α) ≃ Unique α := /-- If `Unique β`, then `Unique α` is equivalent to `α ≃ β`. -/ def uniqueEquivEquivUnique (α : Sort u) (β : Sort v) [Unique β] : Unique α ≃ (α ≃ β) := - equivOfSubsingletonOfSubsingleton (fun _ => Equiv.equivOfUnique _ _) Equiv.unique + equivOfSubsingletonOfSubsingleton (fun _ => Equiv.ofUnique _ _) Equiv.unique namespace Function diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 8246c3649ba3e..c1382615c0f6e 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -398,14 +398,16 @@ def equivEmptyEquiv (α : Sort u) : α ≃ Empty ≃ IsEmpty α := def propEquivPEmpty {p : Prop} (h : ¬p) : p ≃ PEmpty := @equivPEmpty p <| IsEmpty.prop_iff.2 h /-- If both `α` and `β` have a unique element, then `α ≃ β`. -/ -def equivOfUnique (α β : Sort _) [Unique.{u} α] [Unique.{v} β] : α ≃ β where +def ofUnique (α β : Sort _) [Unique.{u} α] [Unique.{v} β] : α ≃ β where toFun := default invFun := default left_inv _ := Subsingleton.elim _ _ right_inv _ := Subsingleton.elim _ _ +@[deprecated (since := "2024-12-26")] alias equivOfUnique := ofUnique + /-- If `α` has a unique element, then it is equivalent to any `PUnit`. -/ -def equivPUnit (α : Sort u) [Unique α] : α ≃ PUnit.{v} := equivOfUnique α _ +def equivPUnit (α : Sort u) [Unique α] : α ≃ PUnit.{v} := ofUnique α _ /-- The `Sort` of proofs of a true proposition is equivalent to `PUnit`. -/ def propEquivPUnit {p : Prop} (h : p) : p ≃ PUnit.{0} := @equivPUnit p <| uniqueProp h diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean index 12eef5dd443c3..d597d44b6f446 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean @@ -357,7 +357,7 @@ protected theorem measurable_comp_iff {f : β → γ} (e : α ≃ᵐ β) : /-- Any two types with unique elements are measurably equivalent. -/ def ofUniqueOfUnique (α β : Type*) [MeasurableSpace α] [MeasurableSpace β] [Unique α] [Unique β] : α ≃ᵐ β where - toEquiv := equivOfUnique α β + toEquiv := ofUnique α β measurable_toFun := Subsingleton.measurable measurable_invFun := Subsingleton.measurable diff --git a/Mathlib/Order/Estimator.lean b/Mathlib/Order/Estimator.lean index c943ad741ebe7..404ab52dbd6a8 100644 --- a/Mathlib/Order/Estimator.lean +++ b/Mathlib/Order/Estimator.lean @@ -77,7 +77,7 @@ instance : WellFoundedGT Unit where wf := ⟨fun .unit => ⟨Unit.unit, nofun⟩⟩ instance (a : α) : WellFoundedGT (Estimator.trivial a) := - let f : Estimator.trivial a ≃o Unit := RelIso.relIsoOfUniqueOfRefl _ _ + let f : Estimator.trivial a ≃o Unit := RelIso.ofUniqueOfRefl _ _ let f' : Estimator.trivial a ↪o Unit := f.toOrderEmbedding f'.wellFoundedGT diff --git a/Mathlib/Order/RelIso/Basic.lean b/Mathlib/Order/RelIso/Basic.lean index cc0083a0b2676..d5c77ea6e959e 100644 --- a/Mathlib/Order/RelIso/Basic.lean +++ b/Mathlib/Order/RelIso/Basic.lean @@ -724,14 +724,18 @@ def relIsoOfIsEmpty (r : α → α → Prop) (s : β → β → Prop) [IsEmpty ⟨Equiv.equivOfIsEmpty α β, @fun a => isEmptyElim a⟩ /-- Two irreflexive relations on a unique type are isomorphic. -/ -def relIsoOfUniqueOfIrrefl (r : α → α → Prop) (s : β → β → Prop) [IsIrrefl α r] +def ofUniqueOfIrrefl (r : α → α → Prop) (s : β → β → Prop) [IsIrrefl α r] [IsIrrefl β s] [Unique α] [Unique β] : r ≃r s := - ⟨Equiv.equivOfUnique α β, iff_of_false (not_rel_of_subsingleton s _ _) + ⟨Equiv.ofUnique α β, iff_of_false (not_rel_of_subsingleton s _ _) (not_rel_of_subsingleton r _ _) ⟩ +@[deprecated (since := "2024-12-26")] alias relIsoOfUniqueOfIrrefl := ofUniqueOfIrrefl + /-- Two reflexive relations on a unique type are isomorphic. -/ -def relIsoOfUniqueOfRefl (r : α → α → Prop) (s : β → β → Prop) [IsRefl α r] [IsRefl β s] +def ofUniqueOfRefl (r : α → α → Prop) (s : β → β → Prop) [IsRefl α r] [IsRefl β s] [Unique α] [Unique β] : r ≃r s := - ⟨Equiv.equivOfUnique α β, iff_of_true (rel_of_subsingleton s _ _) (rel_of_subsingleton r _ _)⟩ + ⟨Equiv.ofUnique α β, iff_of_true (rel_of_subsingleton s _ _) (rel_of_subsingleton r _ _)⟩ + +@[deprecated (since := "2024-12-26")] alias relIsoOfUniqueOfRefl := ofUniqueOfRefl end RelIso diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean index fedc31c3c9a40..b53ab4e4d4058 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean @@ -98,7 +98,7 @@ def actionDiagonalSucc (G : Type u) [Group G] : ∀ n : ℕ, diagonal G (n + 1) ≅ leftRegular G ⊗ Action.mk (Fin n → G) 1 | 0 => diagonalOneIsoLeftRegular G ≪≫ - (ρ_ _).symm ≪≫ tensorIso (Iso.refl _) (tensorUnitIso (Equiv.equivOfUnique PUnit _).toIso) + (ρ_ _).symm ≪≫ tensorIso (Iso.refl _) (tensorUnitIso (Equiv.ofUnique PUnit _).toIso) | n + 1 => diagonalSucc _ _ ≪≫ tensorIso (Iso.refl _) (actionDiagonalSucc G n) ≪≫ diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 72adc4a40aa06..21c7cf5ee7db4 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -633,8 +633,7 @@ instance Ideal.uniqueFactorizationMonoid : UniqueFactorizationMonoid (Ideal A) : ⟨x * y, Ideal.mul_mem_mul x_mem y_mem, mt this.isPrime.mem_or_mem (not_or_intro x_not_mem y_not_mem)⟩⟩, Prime.irreducible⟩ } -instance Ideal.normalizationMonoid : NormalizationMonoid (Ideal A) := - normalizationMonoidOfUniqueUnits +instance Ideal.normalizationMonoid : NormalizationMonoid (Ideal A) := .ofUniqueUnits @[simp] theorem Ideal.dvd_span_singleton {I : Ideal A} {x : A} : I ∣ Ideal.span {x} ↔ x ∈ I := diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 0761c27a15594..d2919d4a8a761 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -440,7 +440,7 @@ instance : Nontrivial Cardinal.{u} := ⟨⟨1, 0, mk_ne_zero _⟩⟩ theorem mk_eq_one (α : Type u) [Subsingleton α] [Nonempty α] : #α = 1 := - let ⟨_⟩ := nonempty_unique α; (Equiv.equivOfUnique α (ULift (Fin 1))).cardinal_eq + let ⟨_⟩ := nonempty_unique α; (Equiv.ofUnique α (ULift (Fin 1))).cardinal_eq theorem le_one_iff_subsingleton {α : Type u} : #α ≤ 1 ↔ Subsingleton α := ⟨fun ⟨f⟩ => ⟨fun _ _ => f.injective (Subsingleton.elim _ _)⟩, fun ⟨h⟩ => diff --git a/Mathlib/SetTheory/Game/Nim.lean b/Mathlib/SetTheory/Game/Nim.lean index dba76e6378507..367f316e81fa6 100644 --- a/Mathlib/SetTheory/Game/Nim.lean +++ b/Mathlib/SetTheory/Game/Nim.lean @@ -166,7 +166,7 @@ theorem nim_one_moveRight (x) : (nim 1).moveRight x = nim 0 := by simp def nimOneRelabelling : nim 1 ≡r star := by rw [nim_def] refine ⟨?_, ?_, fun i => ?_, fun j => ?_⟩ - any_goals dsimp; apply Equiv.equivOfUnique + any_goals dsimp; apply Equiv.ofUnique all_goals simpa [enumIsoToType] using nimZeroRelabelling theorem nim_one_equiv : nim 1 ≈ star := diff --git a/Mathlib/SetTheory/Game/Ordinal.lean b/Mathlib/SetTheory/Game/Ordinal.lean index cf5a7e8f4cf97..cca43ad7f0086 100644 --- a/Mathlib/SetTheory/Game/Ordinal.lean +++ b/Mathlib/SetTheory/Game/Ordinal.lean @@ -102,7 +102,7 @@ theorem one_toPGame_moveLeft (x) : (toPGame 1).moveLeft x = toPGame 0 := by simp /-- `1.toPGame` has the same moves as `1`. -/ noncomputable def oneToPGameRelabelling : toPGame 1 ≡r 1 := - ⟨Equiv.equivOfUnique _ _, Equiv.equivOfIsEmpty _ _, fun i => by + ⟨Equiv.ofUnique _ _, Equiv.equivOfIsEmpty _ _, fun i => by simpa using zeroToPGameRelabelling, isEmptyElim⟩ theorem toPGame_one : toPGame 1 ≈ 1 := diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 4b873fb5d2664..a663b8dadc323 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -192,7 +192,7 @@ theorem type_empty : type (@EmptyRelation Empty) = 0 := theorem type_eq_one_of_unique (r) [IsWellOrder α r] [Nonempty α] [Subsingleton α] : type r = 1 := by cases nonempty_unique α - exact (RelIso.relIsoOfUniqueOfIrrefl r _).ordinal_type_eq + exact (RelIso.ofUniqueOfIrrefl r _).ordinal_type_eq @[simp] theorem type_eq_one_iff_unique [IsWellOrder α r] : type r = 1 ↔ Nonempty (Unique α) := diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index e422388b09bca..f6859612b043c 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -667,7 +667,7 @@ def punitProd : PUnit × X ≃ₜ X := /-- If both `X` and `Y` have a unique element, then `X ≃ₜ Y`. -/ @[simps!] def homeomorphOfUnique [Unique X] [Unique Y] : X ≃ₜ Y := - { Equiv.equivOfUnique X Y with + { Equiv.ofUnique X Y with continuous_toFun := continuous_const continuous_invFun := continuous_const } From 2a7dd2dfedddabf87dfb520cf6cf28a888439438 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 27 Dec 2024 12:17:51 +0000 Subject: [PATCH 010/189] feat(MetricSpace): add `Metric.biInter_gt_ball` etc (#20262) --- Mathlib/Order/Basic.lean | 7 +++++++ .../Topology/MetricSpace/Pseudo/Lemmas.lean | 18 ++++++++++++++++++ 2 files changed, 25 insertions(+) diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index a76d787182917..7649faffca083 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -1281,6 +1281,13 @@ theorem eq_of_le_of_forall_ge_of_dense [LinearOrder α] [DenselyOrdered α] {a (h₂ : ∀ a₃ < a₁, a₃ ≤ a₂) : a₁ = a₂ := (le_of_forall_ge_of_dense h₂).antisymm h₁ +theorem forall_lt_le_iff [LinearOrder α] [DenselyOrdered α] {a b : α} : (∀ c < a, c ≤ b) ↔ a ≤ b := + ⟨le_of_forall_ge_of_dense, fun hab _c hca ↦ hca.le.trans hab⟩ + +theorem forall_gt_ge_iff [LinearOrder α] [DenselyOrdered α] {a b : α} : + (∀ c, a < c → b ≤ c) ↔ b ≤ a := + forall_lt_le_iff (α := αᵒᵈ) + theorem dense_or_discrete [LinearOrder α] (a₁ a₂ : α) : (∃ a, a₁ < a ∧ a < a₂) ∨ (∀ a, a₁ < a → a₂ ≤ a) ∧ ∀ a < a₂, a ≤ a₁ := or_iff_not_imp_left.2 fun h ↦ diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean b/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean index 2a58189df2473..6b6cd86da53a2 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean @@ -89,4 +89,22 @@ lemma exists_isCompact_closedBall [WeaklyLocallyCompactSpace α] (x : α) : eventually_nhdsWithin_of_eventually_nhds (eventually_isCompact_closedBall x) simpa only [and_comm] using (this.and self_mem_nhdsWithin).exists +theorem biInter_gt_closedBall (x : α) (r : ℝ) : ⋂ r' > r, closedBall x r' = closedBall x r := by + ext + simp [forall_gt_ge_iff] + +theorem biInter_gt_ball (x : α) (r : ℝ) : ⋂ r' > r, ball x r' = closedBall x r := by + ext + simp [forall_lt_iff_le'] + +theorem biUnion_lt_ball (x : α) (r : ℝ) : ⋃ r' < r, ball x r' = ball x r := by + ext + rw [← not_iff_not] + simp [forall_lt_le_iff] + +theorem biUnion_lt_closedBall (x : α) (r : ℝ) : ⋃ r' < r, closedBall x r' = ball x r := by + ext + rw [← not_iff_not] + simp [forall_lt_iff_le] + end Metric From cdf32a02ac52bf67f2540f3aa9a1c74ff34e5ab4 Mon Sep 17 00:00:00 2001 From: smorel394 Date: Fri, 27 Dec 2024 13:01:13 +0000 Subject: [PATCH 011/189] feat(CategoryTheory/Adjunction/Additive): adjunctions between additive functors (#20083) This provides some results and constructions for adjunctions between functors on preadditive categories: * If one of the adjoint functors is additive, so is the other. * If one of the adjoint functors is additive, the equivalence `Adjunction.homEquiv` lifts to an additive equivalence `Adjunction.homAddEquiv`. * We also give a version of this additive equivalence as an isomorphism of `preadditiveYoneda` functors (analogous to `Adjunction.compYonedaIso`), in `Adjunction.compPreadditiveYonedaIso`. Co-authored-by: smorel394 <67864981+smorel394@users.noreply.github.com> --- Mathlib.lean | 1 + .../CategoryTheory/Adjunction/Additive.lean | 131 ++++++++++++++++++ 2 files changed, 132 insertions(+) create mode 100644 Mathlib/CategoryTheory/Adjunction/Additive.lean diff --git a/Mathlib.lean b/Mathlib.lean index c7c011d73f350..a5276dded855a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1612,6 +1612,7 @@ import Mathlib.CategoryTheory.Action.Continuous import Mathlib.CategoryTheory.Action.Limits import Mathlib.CategoryTheory.Action.Monoidal import Mathlib.CategoryTheory.Adhesive +import Mathlib.CategoryTheory.Adjunction.Additive import Mathlib.CategoryTheory.Adjunction.AdjointFunctorTheorems import Mathlib.CategoryTheory.Adjunction.Basic import Mathlib.CategoryTheory.Adjunction.Comma diff --git a/Mathlib/CategoryTheory/Adjunction/Additive.lean b/Mathlib/CategoryTheory/Adjunction/Additive.lean new file mode 100644 index 0000000000000..7367e284b4350 --- /dev/null +++ b/Mathlib/CategoryTheory/Adjunction/Additive.lean @@ -0,0 +1,131 @@ +/- +Copyright (c) 2024 Sophie Morel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sophie Morel, Joël Riou +-/ +import Mathlib.CategoryTheory.Preadditive.Yoneda.Basic + +/-! +# Adjunctions between additive functors. + +This provides some results and constructions for adjunctions between functors on +preadditive categories: +* If one of the adjoint functors is additive, so is the other. +* If one of the adjoint functors is additive, the equivalence `Adjunction.homEquiv` lifts to +an additive equivalence `Adjunction.homAddEquiv`. +* We also give a version of this additive equivalence as an isomorphism of `preadditiveYoneda` +functors (analogous to `Adjunction.compYonedaIso`), in `Adjunction.compPreadditiveYonedaIso`. + +-/ + +universe u₁ u₂ v₁ v₂ + +namespace CategoryTheory + +namespace Adjunction + +open CategoryTheory Category Functor + +variable {C : Type u₁} {D : Type u₂} [Category.{v₁} C] [Category.{v₂} D] [Preadditive C] + [Preadditive D] {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) + +include adj +lemma right_adjoint_additive [F.Additive] : G.Additive where + map_add {X Y} f g := (adj.homEquiv _ _).symm.injective (by simp [homEquiv_counit]) + +lemma left_adjoint_additive [G.Additive] : F.Additive where + map_add {X Y} f g := (adj.homEquiv _ _).injective (by simp [homEquiv_unit]) + +variable [F.Additive] + +/-- If we have an adjunction `adj : F ⊣ G` of functors between preadditive categories, +and if `F` is additive, then the hom set equivalence upgrades to an `AddEquiv`. +Note that `F` is additive if and only if `G` is, by `Adjunction.right_adjoint_additive` and +`Adjunction.left_adjoint_additive`. +-/ +def homAddEquiv (X : C) (Y : D) : AddEquiv (F.obj X ⟶ Y) (X ⟶ G.obj Y) := + { adj.homEquiv _ _ with + map_add' _ _ := by + have := adj.right_adjoint_additive + simp [homEquiv_apply] } + +@[simp] +lemma homAddEquiv_apply (X : C) (Y : D) (f : F.obj X ⟶ Y) : + adj.homAddEquiv X Y f = adj.homEquiv X Y f := rfl + +@[simp] +lemma homAddEquiv_symm_apply (X : C) (Y : D) (f : X ⟶ G.obj Y) : + (adj.homAddEquiv X Y).symm f = (adj.homEquiv X Y).symm f := rfl + +@[simp] +lemma homAddEquiv_zero (X : C) (Y : D) : adj.homEquiv X Y 0 = 0 := map_zero (adj.homAddEquiv X Y) + +@[simp] +lemma homAddEquiv_add (X : C) (Y : D) (f f' : F.obj X ⟶ Y) : + adj.homEquiv X Y (f + f') = adj.homEquiv X Y f + adj.homEquiv X Y f' := + map_add (adj.homAddEquiv X Y) _ _ + +@[simp] +lemma homAddEquiv_sub (X : C) (Y : D) (f f' : F.obj X ⟶ Y) : + adj.homEquiv X Y (f - f') = adj.homEquiv X Y f - adj.homEquiv X Y f' := + map_sub (adj.homAddEquiv X Y) _ _ + +@[simp] +lemma homAddEquiv_neg (X : C) (Y : D) (f : F.obj X ⟶ Y) : + adj.homEquiv X Y (- f) = - adj.homEquiv X Y f := map_neg (adj.homAddEquiv X Y) _ + +@[simp] +lemma homAddEquiv_symm_zero (X : C) (Y : D) : + (adj.homEquiv X Y).symm 0 = 0 := map_zero (adj.homAddEquiv X Y).symm + +@[simp] +lemma homAddEquiv_symm_add (X : C) (Y : D) (f f' : X ⟶ G.obj Y) : + (adj.homEquiv X Y).symm (f + f') = (adj.homEquiv X Y).symm f + (adj.homEquiv X Y).symm f' := + map_add (adj.homAddEquiv X Y).symm _ _ + +@[simp] +lemma homAddEquiv_symm_sub (X : C) (Y : D) (f f' : X ⟶ G.obj Y) : + (adj.homEquiv X Y).symm (f - f') = (adj.homEquiv X Y).symm f - (adj.homEquiv X Y).symm f' := + map_sub (adj.homAddEquiv X Y).symm _ _ + +@[simp] +lemma homAddEquiv_symm_neg (X : C) (Y : D) (f : X ⟶ G.obj Y) : + (adj.homEquiv X Y).symm (- f) = - (adj.homEquiv X Y).symm f := + map_neg (adj.homAddEquiv X Y).symm _ + +open Opposite in +/-- If we have an adjunction `adj : F ⊣ G` of functors between preadditive categories, +and if `F` is additive, then the hom set equivalence upgrades to an isomorphism between +`G ⋙ preadditiveYoneda` and `preadditiveYoneda ⋙ F`, once we throw in the necessary +universe lifting functors. +Note that `F` is additive if and only if `G` is, by `Adjunction.right_adjoint_additive` and +`Adjunction.left_adjoint_additive`. +-/ +def compPreadditiveYonedaIso : + G ⋙ preadditiveYoneda ⋙ (whiskeringRight _ _ _).obj AddCommGrp.uliftFunctor.{max v₁ v₂} ≅ + preadditiveYoneda ⋙ (whiskeringLeft _ _ _).obj F.op ⋙ + (whiskeringRight _ _ _).obj AddCommGrp.uliftFunctor.{max v₁ v₂} := + NatIso.ofComponents + (fun Y ↦ NatIso.ofComponents + (fun X ↦ (AddEquiv.ulift.trans ((adj.homAddEquiv (unop X) Y).symm.trans + AddEquiv.ulift.symm)).toAddCommGrpIso) + (fun g ↦ by + ext ⟨y⟩ + exact AddEquiv.ulift.injective (adj.homEquiv_naturality_left_symm g.unop y))) + (fun f ↦ by + ext _ ⟨x⟩ + exact AddEquiv.ulift.injective ((adj.homEquiv_naturality_right_symm x f))) + +lemma compPreadditiveYonedaIso_hom_app_app_apply (X : Cᵒᵖ) (Y : D) + (a : ULift.{max v₁ v₂, v₁} (Opposite.unop X ⟶ G.obj Y)) : + ((adj.compPreadditiveYonedaIso.hom.app Y).app X) a = + ULift.up ((adj.homEquiv (Opposite.unop X) Y).symm (AddEquiv.ulift a)) := rfl + +lemma compPreadditiveYonedaIso_inv_app_app_apply (X : Cᵒᵖ) (Y : D) + (a : ULift.{max v₁ v₂, v₂} (F.obj (Opposite.unop X) ⟶ Y)) : + ((adj.compPreadditiveYonedaIso.inv.app Y).app X) a = + ULift.up ((adj.homEquiv (Opposite.unop X) Y) (AddEquiv.ulift a)) := rfl + +end Adjunction + +end CategoryTheory From e85c35e9579aa055568b5d7e05c936a1f4e992fd Mon Sep 17 00:00:00 2001 From: Seewoo Lee Date: Fri, 27 Dec 2024 18:42:23 +0000 Subject: [PATCH 012/189] chore: protect `Polynomial.map_eq_zero` (#20268) Protect `Polynomial.map_eq_zero` to avoid conflict with `_root_.map_eq_zero`. --- Mathlib/Algebra/Polynomial/FieldDivision.lean | 10 +++++----- Mathlib/Algebra/Polynomial/Splits.lean | 2 +- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/FieldDivision.lean b/Mathlib/Algebra/Polynomial/FieldDivision.lean index 3b5ed5bdcc23f..e5294cbb83106 100644 --- a/Mathlib/Algebra/Polynomial/FieldDivision.lean +++ b/Mathlib/Algebra/Polynomial/FieldDivision.lean @@ -255,13 +255,13 @@ theorem degree_pos_of_ne_zero_of_nonunit (hp0 : p ≠ 0) (hp : ¬IsUnit p) : 0 < exact hp (IsUnit.map C (IsUnit.mk0 (coeff p 0) (mt C_inj.2 (by simpa using hp0)))) @[simp] -theorem map_eq_zero [Semiring S] [Nontrivial S] (f : R →+* S) : p.map f = 0 ↔ p = 0 := by +protected theorem map_eq_zero [Semiring S] [Nontrivial S] (f : R →+* S) : p.map f = 0 ↔ p = 0 := by simp only [Polynomial.ext_iff] congr! simp [map_eq_zero, coeff_map, coeff_zero] theorem map_ne_zero [Semiring S] [Nontrivial S] {f : R →+* S} (hp : p ≠ 0) : p.map f ≠ 0 := - mt (map_eq_zero f).1 hp + mt (Polynomial.map_eq_zero f).1 hp @[simp] theorem degree_map [Semiring S] [Nontrivial S] (p : R[X]) (f : R →+* S) : @@ -494,7 +494,7 @@ theorem rootSet_prod [CommRing S] [IsDomain S] [Algebra R S] {ι : Type*} (f : classical simp only [rootSet, aroots, ← Finset.mem_coe] rw [Polynomial.map_prod, roots_prod, Finset.bind_toFinset, s.val_toFinset, Finset.coe_biUnion] - rwa [← Polynomial.map_prod, Ne, map_eq_zero] + rwa [← Polynomial.map_prod, Ne, Polynomial.map_eq_zero] theorem exists_root_of_degree_eq_one (h : degree p = 1) : ∃ x, IsRoot p x := ⟨-(p.coeff 0 / p.coeff 1), by @@ -555,10 +555,10 @@ theorem coe_normUnit_of_ne_zero [DecidableEq R] (hp : p ≠ 0) : theorem map_dvd_map' [Field k] (f : R →+* k) {x y : R[X]} : x.map f ∣ y.map f ↔ x ∣ y := by by_cases H : x = 0 - · rw [H, Polynomial.map_zero, zero_dvd_iff, zero_dvd_iff, map_eq_zero] + · rw [H, Polynomial.map_zero, zero_dvd_iff, zero_dvd_iff, Polynomial.map_eq_zero] · classical rw [← normalize_dvd_iff, ← @normalize_dvd_iff R[X], normalize_apply, normalize_apply, - coe_normUnit_of_ne_zero H, coe_normUnit_of_ne_zero (mt (map_eq_zero f).1 H), + coe_normUnit_of_ne_zero H, coe_normUnit_of_ne_zero (mt (Polynomial.map_eq_zero f).1 H), leadingCoeff_map, ← map_inv₀ f, ← map_C, ← Polynomial.map_mul, map_dvd_map _ f.injective (monic_mul_leadingCoeff_inv H)] diff --git a/Mathlib/Algebra/Polynomial/Splits.lean b/Mathlib/Algebra/Polynomial/Splits.lean index 49d517483faae..faffbfa26932e 100644 --- a/Mathlib/Algebra/Polynomial/Splits.lean +++ b/Mathlib/Algebra/Polynomial/Splits.lean @@ -263,7 +263,7 @@ variable (i : K →+* L) /-- This lemma is for polynomials over a field. -/ theorem splits_iff (f : K[X]) : Splits i f ↔ f = 0 ∨ ∀ {g : L[X]}, Irreducible g → g ∣ f.map i → degree g = 1 := by - rw [Splits, map_eq_zero] + rw [Splits, Polynomial.map_eq_zero] /-- This lemma is for polynomials over a field. -/ theorem Splits.def {i : K →+* L} {f : K[X]} (h : Splits i f) : From f332af85a83e12e1530c438ba2e8b0ed865b0369 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Fri, 27 Dec 2024 19:18:07 +0000 Subject: [PATCH 013/189] chore(RingTheory/LocalRing): strengthen `Module.free_of_flat_of_isLocalRing` (#19846) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Make the fact that the proof works with any `k`-basis of `k ⊗[R] M` explicit and deduce from this that any generating family contains a basis. --- Mathlib/LinearAlgebra/LinearIndependent.lean | 17 +++ Mathlib/RingTheory/LocalRing/Module.lean | 137 ++++++++++++------- 2 files changed, 106 insertions(+), 48 deletions(-) diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index 5472b6646d941..071d3990c70ae 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -1355,6 +1355,23 @@ theorem exists_linearIndependent : exists_linearIndependent_extension (linearIndependent_empty K V) (Set.empty_subset t) exact ⟨b, hb₁, (span_eq_of_le _ hb₂ (Submodule.span_mono hb₁)).symm, hb₃⟩ +/-- Indexed version of `exists_linearIndependent`. -/ +lemma exists_linearIndependent' (v : ι → V) : + ∃ (κ : Type u') (a : κ → ι), Function.Injective a ∧ + Submodule.span K (Set.range (v ∘ a)) = Submodule.span K (Set.range v) ∧ + LinearIndependent K (v ∘ a) := by + obtain ⟨t, ht, hsp, hli⟩ := exists_linearIndependent K (Set.range v) + choose f hf using ht + let s : Set ι := Set.range (fun a : t ↦ f a.property) + have hs {i : ι} (hi : i ∈ s) : v i ∈ t := by obtain ⟨a, rfl⟩ := hi; simp [hf] + let f' (a : s) : t := ⟨v a.val, hs a.property⟩ + refine ⟨s, Subtype.val, Subtype.val_injective, hsp.symm ▸ by congr; aesop, ?_⟩ + · rw [← show Subtype.val ∘ f' = v ∘ Subtype.val by ext; simp] + apply hli.comp + rintro ⟨i, x, rfl⟩ ⟨j, y, rfl⟩ hij + simp only [Subtype.ext_iff, hf] at hij + simp [hij] + variable {K t} /-- `LinearIndependent.extend` adds vectors to a linear independent set `s ⊆ t` until it spans diff --git a/Mathlib/RingTheory/LocalRing/Module.lean b/Mathlib/RingTheory/LocalRing/Module.lean index 40aa2ebe664c3..8aec28a9bca3d 100644 --- a/Mathlib/RingTheory/LocalRing/Module.lean +++ b/Mathlib/RingTheory/LocalRing/Module.lean @@ -28,6 +28,8 @@ 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. -/ +universe u + variable {R} [CommRing R] section @@ -153,61 +155,100 @@ namespace Module variable [IsLocalRing R] -/-- -If `M` is a finitely presented module over a local ring `(R, 𝔪)` such that `m ⊗ M → M` is -injective, then `M` is free. --/ -theorem free_of_maximalIdeal_rTensor_injective [Module.FinitePresentation R M] +/-- If `M` is of finite presentation over a local ring `(R, 𝔪, k)` such that +`𝔪 ⊗ M → M` is injective, then every family of elements that is a `k`-basis of +`k ⊗ M` is an `R`-basis of `M`. -/ +lemma exists_basis_of_basis_baseChange [Module.FinitePresentation R M] + {ι : Type u} (v : ι → M) (hli : LinearIndependent k (TensorProduct.mk R k M 1 ∘ v)) + (hsp : Submodule.span k (Set.range (TensorProduct.mk R k M 1 ∘ v)) = ⊤) (H : Function.Injective ((𝔪).subtype.rTensor M)) : - Module.Free R M := by - let I := Module.Free.ChooseBasisIndex k (k ⊗[R] M) - -- Let `b : I → k ⊗ M` be a `k`-basis. - let b : Basis I k _ := Module.Free.chooseBasis k (k ⊗[R] M) - letI : IsNoetherian k (k ⊗[R] (I →₀ R)) := - isNoetherian_of_isNoetherianRing_of_finite k (k ⊗[R] (I →₀ R)) - choose f hf using TensorProduct.mk_surjective R M k Ideal.Quotient.mk_surjective - -- By choosing an arbitrary lift of `b` to `I → M`, we get a surjection `i : Rᴵ → M`. - let i := Finsupp.linearCombination R (f ∘ b) + ∃ (b : Basis ι R M), ∀ i, b i = v i := by + let bk : Basis ι k (k ⊗[R] M) := Basis.mk hli (by rw [hsp]) + haveI : Finite ι := Module.Finite.finite_basis bk + letI : Fintype ι := Fintype.ofFinite ι + letI : IsNoetherian k (k ⊗[R] (ι →₀ R)) := + isNoetherian_of_isNoetherianRing_of_finite k (k ⊗[R] (ι →₀ R)) + let i := Finsupp.linearCombination R v have hi : Surjective i := by rw [← LinearMap.range_eq_top, Finsupp.range_linearCombination] - exact IsLocalRing.span_eq_top_of_tmul_eq_basis (R := R) (f := f ∘ b) b (fun _ ↦ hf _) + refine IsLocalRing.span_eq_top_of_tmul_eq_basis (R := R) (f := v) bk + (fun _ ↦ by simp [bk]) have : Module.Finite R (LinearMap.ker i) := by constructor exact (Submodule.fg_top _).mpr (Module.FinitePresentation.fg_ker i hi) -- We claim that `i` is actually a bijection, - -- hence an isomorphism exhibing `M` as the free module `Rᴵ` - refine Module.Free.of_equiv (LinearEquiv.ofBijective i ⟨?_, hi⟩) - -- By Nakayama's lemma, it suffices to show that `k ⊗ ker(i) = 0`. - rw [← LinearMap.ker_eq_bot, ← Submodule.subsingleton_iff_eq_bot, - ← IsLocalRing.subsingleton_tensorProduct (R := R), subsingleton_iff_forall_eq 0] - have : Function.Surjective (i.baseChange k) := i.lTensor_surjective _ hi - -- By construction, `k ⊗ i : kᴵ → k ⊗ M` is bijective. - have hi' : Function.Bijective (i.baseChange k) := by - refine ⟨?_, this⟩ - rw [← LinearMap.ker_eq_bot (M := k ⊗[R] (I →₀ R)) (f := i.baseChange k), - ← Submodule.finrank_eq_zero (R := k) (M := k ⊗[R] (I →₀ R)), - ← Nat.add_right_inj (n := Module.finrank k (LinearMap.range <| i.baseChange k)), - LinearMap.finrank_range_add_finrank_ker (V := k ⊗[R] (I →₀ R)), - LinearMap.range_eq_top.mpr this, finrank_top] - simp only [Module.finrank_tensorProduct, Module.finrank_self, - Module.finrank_finsupp_self, one_mul, add_zero] - rw [Module.finrank_eq_card_chooseBasisIndex] - -- On the other hand, `m ⊗ M → M` injective => `Tor₁(k, M) = 0` => `k ⊗ ker(i) → kᴵ` injective. - intro x - refine lTensor_injective_of_exact_of_exact_of_rTensor_injective - (N₁ := LinearMap.ker i) (N₂ := I →₀ R) (N₃ := M) - (f₁ := (𝔪).subtype) (f₂ := Submodule.mkQ 𝔪) - (g₁ := (LinearMap.ker i).subtype) (g₂ := i) (LinearMap.exact_subtype_mkQ 𝔪) - (Submodule.mkQ_surjective _) (LinearMap.exact_subtype_ker_map i) hi H ?_ ?_ - · apply Module.Flat.lTensor_preserves_injective_linearMap - (N := LinearMap.ker i) (N' := I →₀ R) - (L := (LinearMap.ker i).subtype) - exact Subtype.val_injective - · apply hi'.injective - rw [LinearMap.baseChange_eq_ltensor] - erw [← LinearMap.comp_apply (i.lTensor k), ← LinearMap.lTensor_comp] - rw [(LinearMap.exact_subtype_ker_map i).linearMap_comp_eq_zero] - simp only [LinearMap.lTensor_zero, LinearMap.zero_apply, map_zero] + -- hence `v` induces an isomorphism `M ≃[R] Rᴵ` showing that `v` is a basis. + let iequiv : (ι →₀ R) ≃ₗ[R] M := by + refine LinearEquiv.ofBijective i ⟨?_, hi⟩ + -- By Nakayama's lemma, it suffices to show that `k ⊗ ker(i) = 0`. + rw [← LinearMap.ker_eq_bot, ← Submodule.subsingleton_iff_eq_bot, + ← IsLocalRing.subsingleton_tensorProduct (R := R), subsingleton_iff_forall_eq 0] + have : Function.Surjective (i.baseChange k) := i.lTensor_surjective _ hi + -- By construction, `k ⊗ i : kᴵ → k ⊗ M` is bijective. + have hi' : Function.Bijective (i.baseChange k) := by + refine ⟨?_, this⟩ + rw [← LinearMap.ker_eq_bot (M := k ⊗[R] (ι →₀ R)) (f := i.baseChange k), + ← Submodule.finrank_eq_zero (R := k) (M := k ⊗[R] (ι →₀ R)), + ← Nat.add_right_inj (n := Module.finrank k (LinearMap.range <| i.baseChange k)), + LinearMap.finrank_range_add_finrank_ker (V := k ⊗[R] (ι →₀ R)), + LinearMap.range_eq_top.mpr this, finrank_top] + simp only [Module.finrank_tensorProduct, Module.finrank_self, + Module.finrank_finsupp_self, one_mul, add_zero] + rw [Module.finrank_eq_card_basis bk] + -- On the other hand, `m ⊗ M → M` injective => `Tor₁(k, M) = 0` => `k ⊗ ker(i) → kᴵ` injective. + intro x + refine lTensor_injective_of_exact_of_exact_of_rTensor_injective + (N₁ := LinearMap.ker i) (N₂ := ι →₀ R) (N₃ := M) + (f₁ := (𝔪).subtype) (f₂ := Submodule.mkQ 𝔪) + (g₁ := (LinearMap.ker i).subtype) (g₂ := i) (LinearMap.exact_subtype_mkQ 𝔪) + (Submodule.mkQ_surjective _) (LinearMap.exact_subtype_ker_map i) hi H ?_ ?_ + · apply Module.Flat.lTensor_preserves_injective_linearMap + (N := LinearMap.ker i) (N' := ι →₀ R) + (L := (LinearMap.ker i).subtype) + exact Subtype.val_injective + · apply hi'.injective + rw [LinearMap.baseChange_eq_ltensor] + erw [← LinearMap.comp_apply (i.lTensor k), ← LinearMap.lTensor_comp] + rw [(LinearMap.exact_subtype_ker_map i).linearMap_comp_eq_zero] + simp only [LinearMap.lTensor_zero, LinearMap.zero_apply, map_zero] + use Basis.ofRepr iequiv.symm + intro j + simp [iequiv, i] + +/-- +If `M` is a finitely presented module over a local ring `(R, 𝔪)` such that `m ⊗ M → M` is +injective, then every generating family contains a basis. +-/ +lemma exists_basis_of_span_of_maximalIdeal_rTensor_injective [Module.FinitePresentation R M] + (H : Function.Injective ((𝔪).subtype.rTensor M)) + {ι : Type u} (v : ι → M) (hv : Submodule.span R (Set.range v) = ⊤) : + ∃ (κ : Type u) (a : κ → ι) (b : Basis κ R M), ∀ i, b i = v (a i) := by + have := (map_tensorProduct_mk_eq_top (N := Submodule.span R (Set.range v))).mpr hv + rw [← Submodule.span_image, ← Set.range_comp, eq_top_iff, ← SetLike.coe_subset_coe, + Submodule.top_coe] at this + have : Submodule.span k (Set.range (TensorProduct.mk R k M 1 ∘ v)) = ⊤ := by + rw [eq_top_iff] + exact Set.Subset.trans this (Submodule.span_subset_span _ _ _) + obtain ⟨κ, a, ha, hsp, hli⟩ := exists_linearIndependent' k (TensorProduct.mk R k M 1 ∘ v) + rw [this] at hsp + obtain ⟨b, hb⟩ := exists_basis_of_basis_baseChange (v ∘ a) hli hsp H + use κ, a, b, hb + +lemma exists_basis_of_span_of_flat [Module.FinitePresentation R M] [Module.Flat R M] + {ι : Type u} (v : ι → M) (hv : Submodule.span R (Set.range v) = ⊤) : + ∃ (κ : Type u) (a : κ → ι) (b : Basis κ R M), ∀ i, b i = v (a i) := + exists_basis_of_span_of_maximalIdeal_rTensor_injective + (Module.Flat.rTensor_preserves_injective_linearMap (𝔪).subtype Subtype.val_injective) v hv + +/-- +If `M` is a finitely presented module over a local ring `(R, 𝔪)` such that `m ⊗ M → M` is +injective, then `M` is free. +-/ +theorem free_of_maximalIdeal_rTensor_injective [Module.FinitePresentation R M] + (H : Function.Injective ((𝔪).subtype.rTensor M)) : + Module.Free R M := by + obtain ⟨_, _, b, _⟩ := exists_basis_of_span_of_maximalIdeal_rTensor_injective H id (by simp) + exact Free.of_basis b -- TODO: Generalise this to finite free modules. theorem free_of_flat_of_isLocalRing [Module.FinitePresentation R P] [Module.Flat R P] : From 1b71c694536f17c6bf6936100339f1cdad203588 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 27 Dec 2024 19:43:12 +0000 Subject: [PATCH 014/189] chore: move Logic/Small/{Group,Ring,Module} to Algebra/Small (#20123) This commit establishes that no file in `Logic/` imports any file in `Algebra/`. --- Mathlib.lean | 6 +++--- Mathlib/Algebra/Category/MonCat/Limits.lean | 2 +- Mathlib/{Logic => Algebra}/Small/Group.lean | 0 Mathlib/{Logic => Algebra}/Small/Module.lean | 4 ++-- Mathlib/{Logic => Algebra}/Small/Ring.lean | 0 5 files changed, 6 insertions(+), 6 deletions(-) rename Mathlib/{Logic => Algebra}/Small/Group.lean (100%) rename Mathlib/{Logic => Algebra}/Small/Module.lean (93%) rename Mathlib/{Logic => Algebra}/Small/Ring.lean (100%) diff --git a/Mathlib.lean b/Mathlib.lean index a5276dded855a..19697cf5f9461 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -919,6 +919,9 @@ import Mathlib.Algebra.Ring.WithZero import Mathlib.Algebra.RingQuot import Mathlib.Algebra.SMulWithZero import Mathlib.Algebra.SkewMonoidAlgebra.Basic +import Mathlib.Algebra.Small.Group +import Mathlib.Algebra.Small.Module +import Mathlib.Algebra.Small.Ring import Mathlib.Algebra.Squarefree.Basic import Mathlib.Algebra.Star.Basic import Mathlib.Algebra.Star.BigOperators @@ -3601,10 +3604,7 @@ import Mathlib.Logic.Relation import Mathlib.Logic.Relator import Mathlib.Logic.Small.Basic import Mathlib.Logic.Small.Defs -import Mathlib.Logic.Small.Group import Mathlib.Logic.Small.List -import Mathlib.Logic.Small.Module -import Mathlib.Logic.Small.Ring import Mathlib.Logic.Small.Set import Mathlib.Logic.Unique import Mathlib.Logic.UnivLE diff --git a/Mathlib/Algebra/Category/MonCat/Limits.lean b/Mathlib/Algebra/Category/MonCat/Limits.lean index ef95015ba092d..9c9d1b5433546 100644 --- a/Mathlib/Algebra/Category/MonCat/Limits.lean +++ b/Mathlib/Algebra/Category/MonCat/Limits.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Mathlib.Algebra.Category.MonCat.Basic +import Mathlib.Algebra.Small.Group import Mathlib.CategoryTheory.Limits.Creates import Mathlib.CategoryTheory.Limits.Types -import Mathlib.Logic.Small.Group /-! # The category of (commutative) (additive) monoids has all limits diff --git a/Mathlib/Logic/Small/Group.lean b/Mathlib/Algebra/Small/Group.lean similarity index 100% rename from Mathlib/Logic/Small/Group.lean rename to Mathlib/Algebra/Small/Group.lean diff --git a/Mathlib/Logic/Small/Module.lean b/Mathlib/Algebra/Small/Module.lean similarity index 93% rename from Mathlib/Logic/Small/Module.lean rename to Mathlib/Algebra/Small/Module.lean index 8dd1444204daa..cb994eb56b817 100644 --- a/Mathlib/Logic/Small/Module.lean +++ b/Mathlib/Algebra/Small/Module.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Mathlib.Logic.Small.Group -import Mathlib.Logic.Small.Ring +import Mathlib.Algebra.Small.Group +import Mathlib.Algebra.Small.Ring /-! # Transfer module and algebra structures from `α` to `Shrink α`. diff --git a/Mathlib/Logic/Small/Ring.lean b/Mathlib/Algebra/Small/Ring.lean similarity index 100% rename from Mathlib/Logic/Small/Ring.lean rename to Mathlib/Algebra/Small/Ring.lean From 054515ea2a5d7f68512934b27548725e733e8cbe Mon Sep 17 00:00:00 2001 From: "Tristan F.-R." Date: Sat, 28 Dec 2024 02:37:52 +0000 Subject: [PATCH 015/189] docs: fix links from docs to test (#20276) Update links that reference the old `test` folder to use `MathlibTest`. Related to #13328. Co-authored-by: Tristan F. --- Mathlib/Analysis/SpecialFunctions/Integrals.lean | 1 - Mathlib/Data/Fin/VecNotation.lean | 2 +- Mathlib/Data/Matrix/Notation.lean | 2 +- Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean | 2 +- Mathlib/Tactic/Bound.lean | 2 +- Mathlib/Tactic/MoveAdd.lean | 2 +- Mathlib/Tactic/Use.lean | 2 +- Mathlib/Tactic/Widget/StringDiagram.lean | 2 +- 8 files changed, 7 insertions(+), 8 deletions(-) diff --git a/Mathlib/Analysis/SpecialFunctions/Integrals.lean b/Mathlib/Analysis/SpecialFunctions/Integrals.lean index b3b78b374fb6f..f8625c6a2ebeb 100644 --- a/Mathlib/Analysis/SpecialFunctions/Integrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/Integrals.lean @@ -21,7 +21,6 @@ This file contains proofs of the integrals of various specific functions. This i * Integrals of the form `sin x ^ m * cos x ^ n` With these lemmas, many simple integrals can be computed by `simp` or `norm_num`. -See `test/integration.lean` for specific examples. This file also contains some facts about the interval integrability of specific functions. diff --git a/Mathlib/Data/Fin/VecNotation.lean b/Mathlib/Data/Fin/VecNotation.lean index aeb31bb6c4a94..9e242fbf4e9e3 100644 --- a/Mathlib/Data/Fin/VecNotation.lean +++ b/Mathlib/Data/Fin/VecNotation.lean @@ -31,7 +31,7 @@ The main new notation is `![a, b]`, which gets expanded to `vecCons a (vecCons b ## Examples -Examples of usage can be found in the `test/matrix.lean` file. +Examples of usage can be found in the `MathlibTest/matrix.lean` file. -/ diff --git a/Mathlib/Data/Matrix/Notation.lean b/Mathlib/Data/Matrix/Notation.lean index a150ccd3763b4..26667c8ae47fb 100644 --- a/Mathlib/Data/Matrix/Notation.lean +++ b/Mathlib/Data/Matrix/Notation.lean @@ -34,7 +34,7 @@ This file provide notation `!![a, b; c, d]` for matrices, which corresponds to ## Examples -Examples of usage can be found in the `test/matrix.lean` file. +Examples of usage can be found in the `MathlibTest/matrix.lean` file. -/ namespace Matrix diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index 230982afbb29a..57b1ced0fed82 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -32,7 +32,7 @@ This file defines the determinant of a matrix, `Matrix.det`, and its essential p ## Implementation notes It is possible to configure `simp` to compute determinants. See the file -`test/matrix.lean` for some examples. +`MathlibTest/matrix.lean` for some examples. -/ diff --git a/Mathlib/Tactic/Bound.lean b/Mathlib/Tactic/Bound.lean index 58f217d6fa30f..8c42262ecda73 100644 --- a/Mathlib/Tactic/Bound.lean +++ b/Mathlib/Tactic/Bound.lean @@ -27,7 +27,7 @@ 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. +See `MathlibTest/Bound/bound.lean` for tests. ### Calc usage diff --git a/Mathlib/Tactic/MoveAdd.lean b/Mathlib/Tactic/MoveAdd.lean index 76c4709fd49a4..5c0f00d0efe32 100644 --- a/Mathlib/Tactic/MoveAdd.lean +++ b/Mathlib/Tactic/MoveAdd.lean @@ -85,7 +85,7 @@ Currently, `move_oper` supports `HAdd.hAdd`, `HMul.hMul`, `And`, `Or`, `Max.max` These lemmas should be added to `Mathlib.MoveAdd.move_oper_simpCtx`. -See `test/MoveAdd.lean` for sample usage of `move_oper`. +See `MathlibTest/MoveAdd.lean` for sample usage of `move_oper`. ## Implementation notes diff --git a/Mathlib/Tactic/Use.lean b/Mathlib/Tactic/Use.lean index 0f48c40b79e55..89551a70b4847 100644 --- a/Mathlib/Tactic/Use.lean +++ b/Mathlib/Tactic/Use.lean @@ -17,7 +17,7 @@ just like the `exists` tactic, but they can be a little more flexible. that more closely matches `use` from mathlib3. Note: The `use!` tactic is almost exactly the mathlib3 `use` except that it does not try -applying `exists_prop`. See the failing test in `test/Use.lean`. +applying `exists_prop`. See the failing test in `MathlibTest/Use.lean`. -/ namespace Mathlib.Tactic diff --git a/Mathlib/Tactic/Widget/StringDiagram.lean b/Mathlib/Tactic/Widget/StringDiagram.lean index 8b6f3e0562b4e..0a6871ad89cdc 100644 --- a/Mathlib/Tactic/Widget/StringDiagram.lean +++ b/Mathlib/Tactic/Widget/StringDiagram.lean @@ -41,7 +41,7 @@ Currently, the string diagram widget provided in this file deals with equalities in monoidal categories. It displays string diagrams corresponding to the morphisms for the left-hand and right-hand sides of the equality. -Some examples can be found in `test/StringDiagram.lean`. +Some examples can be found in `MathlibTest/StringDiagram.lean`. When drawing string diagrams, it is common to ignore associators and unitors. We follow this convention. To do this, we need to extract non-structural morphisms that are not associators From 370858db2a3e9f6fe583ba6c19a5528d5b0f61b8 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Sat, 28 Dec 2024 03:19:57 +0000 Subject: [PATCH 016/189] fix: improve to_additive warning message (#20199) * Also remove an unused simp extension from the test file, but import another low-level without many dependencies. fixes #20181 --- Mathlib/Tactic/ToAdditive/Frontend.lean | 4 +++- MathlibTest/toAdditive.lean | 15 ++++++++++++--- 2 files changed, 15 insertions(+), 4 deletions(-) diff --git a/Mathlib/Tactic/ToAdditive/Frontend.lean b/Mathlib/Tactic/ToAdditive/Frontend.lean index 7d4f031891763..6e1942fba38c6 100644 --- a/Mathlib/Tactic/ToAdditive/Frontend.lean +++ b/Mathlib/Tactic/ToAdditive/Frontend.lean @@ -1183,9 +1183,11 @@ partial def applyAttributes (stx : Syntax) (rawAttrs : Array Syntax) (thisAttr s if linter.existingAttributeWarning.get (← getOptions) then let appliedAttrs ← getAllSimpAttrs src if appliedAttrs.size > 0 then + let appliedAttrs := ", ".intercalate (appliedAttrs.toList.map toString) + -- Note: we're not bothering to print the correct attribute arguments. Linter.logLintIf linter.existingAttributeWarning stx m!"\ The source declaration {src} was given the simp-attribute(s) {appliedAttrs} before \ - calling @[{thisAttr}]. The preferred method is to use \ + calling @[{thisAttr}]. The preferred method is to use something like \ `@[{thisAttr} (attr := {appliedAttrs})]` to apply the attribute to both \ {src} and the target declaration {tgt}." warnAttr stx Lean.Elab.Tactic.Ext.extExtension diff --git a/MathlibTest/toAdditive.lean b/MathlibTest/toAdditive.lean index 91b979c8d8095..a1f60be919017 100644 --- a/MathlibTest/toAdditive.lean +++ b/MathlibTest/toAdditive.lean @@ -1,5 +1,6 @@ import Mathlib.Algebra.Group.Defs import Mathlib.Lean.Exception +import Mathlib.Tactic.ReduceModChar.Ext import Qq.MetaM open Qq Lean Meta Elab Command ToAdditive @@ -109,9 +110,6 @@ lemma foo15 {α β : Type u} [my_has_pow α β] (x : α) (y : β) : foo14 x y = @[to_additive (reorder := 1 2, 4 5) bar16] lemma foo16 {α β : Type u} [my_has_pow α β] (x : α) (y : β) : foo14 x y = (x ^ y) ^ y := foo15 x y -initialize testExt : SimpExtension ← - registerSimpAttr `simp_test "test" - @[to_additive bar17] def foo17 [Group α] (x : α) : α := x * 1 @@ -424,3 +422,14 @@ run_cmd do unless findTranslation? (← getEnv) `localize.r == some `add_localize.r do throwError "1" unless findTranslation? (← getEnv) `localize == some `add_localize do throwError "2" unless findTranslation? (← getEnv) `localize.s == some `add_localize.s do throwError "3" + +/-- +warning: The source declaration one_eq_one was given the simp-attribute(s) reduce_mod_char, simp before calling @[to_additive]. The preferred method is to use something like `@[to_additive (attr := reduce_mod_char, simp)]` to apply the attribute to both one_eq_one and the target declaration zero_eq_zero. +note: this linter can be disabled with `set_option linter.existingAttributeWarning false` +-/ +#guard_msgs in +@[simp, reduce_mod_char, to_additive] +lemma one_eq_one {α : Type*} [One α] : (1 : α) = 1 := rfl + +@[to_additive (attr := reduce_mod_char, simp)] +lemma one_eq_one' {α : Type*} [One α] : (1 : α) = 1 := rfl From 9bfe1e0f9a092c9e74e406ffa0e0f61b84ac51d0 Mon Sep 17 00:00:00 2001 From: LessnessRandomness <56280933+LessnessRandomness@users.noreply.github.com> Date: Sat, 28 Dec 2024 03:19:58 +0000 Subject: [PATCH 017/189] doc: Update Maps.lean (#20207) Added missing apostrophe (G -> G') in the documentation of [SimpleGraph.Embedding](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Combinatorics/SimpleGraph/Maps.html#SimpleGraph.Embedding) --- Mathlib/Combinatorics/SimpleGraph/Maps.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/SimpleGraph/Maps.lean b/Mathlib/Combinatorics/SimpleGraph/Maps.lean index 4fc6a972b04d8..e27bbf65c3096 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Maps.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Maps.lean @@ -203,7 +203,7 @@ abbrev Hom := RelHom G.Adj G'.Adj /-- A graph embedding is an embedding `f` such that for vertices `v w : V`, -`G.Adj (f v) (f w) ↔ G.Adj v w`. Its image is an induced subgraph of G'. +`G'.Adj (f v) (f w) ↔ G.Adj v w`. Its image is an induced subgraph of G'. The notation `G ↪g G'` represents the type of graph embeddings. -/ abbrev Embedding := From aa726920887e2e051a305016c8c6ca2ed50417bb Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 28 Dec 2024 03:20:00 +0000 Subject: [PATCH 018/189] chore: use `app_delab` (#20249) This has the same behavior as `delab app.`, but protects against typos. The feature appeared in https://github.com/leanprover/lean4/pull/4976, though won't be documented until https://github.com/leanprover/lean4/pull/6450. --- Mathlib/Algebra/BigOperators/Expect.lean | 2 +- Mathlib/Algebra/BigOperators/Group/Finset.lean | 4 ++-- Mathlib/Analysis/InnerProductSpace/PiL2.lean | 2 +- Mathlib/CategoryTheory/Closed/Cartesian.lean | 2 +- Mathlib/CategoryTheory/Monoidal/Category.lean | 2 +- Mathlib/Data/Matrix/Notation.lean | 2 +- Mathlib/Data/Prod/Basic.lean | 4 ++-- Mathlib/Data/SetLike/Basic.lean | 2 +- Mathlib/FieldTheory/Adjoin.lean | 2 +- Mathlib/LinearAlgebra/LinearIndependent.lean | 2 +- Mathlib/LinearAlgebra/Matrix/PosDef.lean | 2 +- Mathlib/Order/SetNotation.lean | 8 ++++---- Mathlib/RingTheory/Derivation/DifferentialRing.lean | 2 +- Mathlib/Util/Delaborators.lean | 4 ++-- 14 files changed, 20 insertions(+), 20 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Expect.lean b/Mathlib/Algebra/BigOperators/Expect.lean index df5560dfae390..10df5c9c7e7b3 100644 --- a/Mathlib/Algebra/BigOperators/Expect.lean +++ b/Mathlib/Algebra/BigOperators/Expect.lean @@ -85,7 +85,7 @@ open Batteries.ExtendedBinder /-- Delaborator for `Finset.expect`. The `pp.piBinderTypes` option controls whether to show the domain type when the expect is over `Finset.univ`. -/ -@[scoped delab app.Finset.expect] def delabFinsetExpect : Delab := +@[scoped app_delab Finset.expect] def delabFinsetExpect : Delab := whenPPOption getPPNotation <| withOverApp 6 <| do let #[_, _, _, _, s, f] := (← getExpr).getAppArgs | failure guard <| f.isLambda diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index d8de2d605f7b1..47e77cc8c7ce4 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -218,7 +218,7 @@ open scoped Batteries.ExtendedBinder /-- Delaborator for `Finset.prod`. The `pp.piBinderTypes` option controls whether to show the domain type when the product is over `Finset.univ`. -/ -@[delab app.Finset.prod] def delabFinsetProd : Delab := +@[app_delab Finset.prod] def delabFinsetProd : Delab := whenPPOption getPPNotation <| withOverApp 5 <| do let #[_, _, _, s, f] := (← getExpr).getAppArgs | failure guard <| f.isLambda @@ -239,7 +239,7 @@ to show the domain type when the product is over `Finset.univ`. -/ /-- Delaborator for `Finset.sum`. The `pp.piBinderTypes` option controls whether to show the domain type when the sum is over `Finset.univ`. -/ -@[delab app.Finset.sum] def delabFinsetSum : Delab := +@[app_delab Finset.sum] def delabFinsetSum : Delab := whenPPOption getPPNotation <| withOverApp 5 <| do let #[_, _, _, s, f] := (← getExpr).getAppArgs | failure guard <| f.isLambda diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index 2f76e57c56c2e..e45fd08b5881d 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -119,7 +119,7 @@ macro_rules | `(!$p:subscript[$e:term,*]) => do set_option trace.debug true in /-- Unexpander for the `!₂[x, y, ...]` notation. -/ -@[delab app.DFunLike.coe] +@[app_delab DFunLike.coe] def EuclideanSpace.delabVecNotation : Delab := whenNotPPOption getPPExplicit <| whenPPOption getPPNotation <| withOverApp 6 do -- check that the `(WithLp.equiv _ _).symm` is present diff --git a/Mathlib/CategoryTheory/Closed/Cartesian.lean b/Mathlib/CategoryTheory/Closed/Cartesian.lean index 0607788f02110..c79932f744135 100644 --- a/Mathlib/CategoryTheory/Closed/Cartesian.lean +++ b/Mathlib/CategoryTheory/Closed/Cartesian.lean @@ -117,7 +117,7 @@ notation:20 A " ⟹ " B:19 => (exp A).obj B open Lean PrettyPrinter.Delaborator SubExpr in /-- Delaborator for `Prefunctor.obj` -/ -@[delab app.Prefunctor.obj] +@[app_delab Prefunctor.obj] def delabPrefunctorObjExp : Delab := whenPPOption getPPNotation <| withOverApp 6 <| do let e ← getExpr guard <| e.isAppOfArity' ``Prefunctor.obj 6 diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 59c3f0eca4f26..a0d928049644b 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -124,7 +124,7 @@ scoped notation "𝟙_ " C:max => (MonoidalCategoryStruct.tensorUnit : C) open Lean PrettyPrinter.Delaborator SubExpr in /-- Used to ensure that `𝟙_` notation is used, as the ascription makes this not automatic. -/ -@[delab app.CategoryTheory.MonoidalCategoryStruct.tensorUnit] +@[app_delab CategoryTheory.MonoidalCategoryStruct.tensorUnit] def delabTensorUnit : Delab := whenPPOption getPPNotation <| withOverApp 3 do let e ← getExpr guard <| e.isAppOfArity ``MonoidalCategoryStruct.tensorUnit 3 diff --git a/Mathlib/Data/Matrix/Notation.lean b/Mathlib/Data/Matrix/Notation.lean index 26667c8ae47fb..66a14eeaec9ed 100644 --- a/Mathlib/Data/Matrix/Notation.lean +++ b/Mathlib/Data/Matrix/Notation.lean @@ -111,7 +111,7 @@ macro_rules | `(!![$[,%$commas]*]) => `(@Matrix.of (Fin 0) (Fin $(quote commas.size)) _ ![]) /-- Delaborator for the `!![]` notation. -/ -@[delab app.DFunLike.coe] +@[app_delab DFunLike.coe] def delabMatrixNotation : Delab := whenNotPPOption getPPExplicit <| whenPPOption getPPNotation <| withOverApp 6 do let mkApp3 (.const ``Matrix.of _) (.app (.const ``Fin _) em) (.app (.const ``Fin _) en) _ := diff --git a/Mathlib/Data/Prod/Basic.lean b/Mathlib/Data/Prod/Basic.lean index 98f729aaec3e6..7df6836ea3cac 100644 --- a/Mathlib/Data/Prod/Basic.lean +++ b/Mathlib/Data/Prod/Basic.lean @@ -304,13 +304,13 @@ section delaborators open Lean PrettyPrinter Delaborator /-- Delaborator for `Prod.fst x` as `x.1`. -/ -@[delab app.Prod.fst] +@[app_delab Prod.fst] def delabProdFst : Delab := withOverApp 3 do let x ← SubExpr.withAppArg delab `($(x).1) /-- Delaborator for `Prod.snd x` as `x.2`. -/ -@[delab app.Prod.snd] +@[app_delab Prod.snd] def delabProdSnd : Delab := withOverApp 3 do let x ← SubExpr.withAppArg delab `($(x).2) diff --git a/Mathlib/Data/SetLike/Basic.lean b/Mathlib/Data/SetLike/Basic.lean index 1e120903e60ff..fb74683ca034a 100644 --- a/Mathlib/Data/SetLike/Basic.lean +++ b/Mathlib/Data/SetLike/Basic.lean @@ -117,7 +117,7 @@ open Lean PrettyPrinter.Delaborator SubExpr /-- For terms that match the `CoeSort` instance's body, pretty print as `↥S` rather than as `{ x // x ∈ S }`. The discriminating feature is that membership uses the `SetLike.instMembership` instance. -/ -@[delab app.Subtype] +@[app_delab Subtype] def delabSubtypeSetLike : Delab := whenPPOption getPPNotation do let #[_, .lam n _ body _] := (← getExpr).getAppArgs | failure guard <| body.isAppOf ``Membership.mem diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index 464f52592988b..6c85ffedec641 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -600,7 +600,7 @@ generated by these elements. -/ scoped macro:max K:term "⟮" xs:term,* "⟯" : term => do ``(adjoin $K $(← mkInsertTerm xs.getElems)) open Lean PrettyPrinter.Delaborator SubExpr in -@[delab app.IntermediateField.adjoin] +@[app_delab IntermediateField.adjoin] partial def delabAdjoinNotation : Delab := whenPPOption getPPNotation do let e ← getExpr guard <| e.isAppOfArity ``adjoin 6 diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index 071d3990c70ae..aa5849bf7df41 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -109,7 +109,7 @@ in case the family of vectors is over a `Set`. Type hints look like `LinearIndependent fun (v : ↑s) => ↑v` or `LinearIndependent (ι := ↑s) f`, depending on whether the family is a lambda expression or not. -/ -@[delab app.LinearIndependent] +@[app_delab LinearIndependent] def delabLinearIndependent : Delab := whenPPOption getPPNotation <| whenNotPPOption getPPAnalysisSkip <| diff --git a/Mathlib/LinearAlgebra/Matrix/PosDef.lean b/Mathlib/LinearAlgebra/Matrix/PosDef.lean index 73d3c0be13503..fb373744d8173 100644 --- a/Mathlib/LinearAlgebra/Matrix/PosDef.lean +++ b/Mathlib/LinearAlgebra/Matrix/PosDef.lean @@ -172,7 +172,7 @@ noncomputable def sqrt : Matrix n n 𝕜 := open Lean PrettyPrinter.Delaborator SubExpr in /-- Custom elaborator to produce output like `(_ : PosSemidef A).sqrt` in the goal view. -/ -@[delab app.Matrix.PosSemidef.sqrt] +@[app_delab Matrix.PosSemidef.sqrt] def delabSqrt : Delab := whenPPOption getPPNotation <| whenNotPPOption getPPAnalysisSkip <| diff --git a/Mathlib/Order/SetNotation.lean b/Mathlib/Order/SetNotation.lean index 7feb44d478558..fb19f19f22d55 100644 --- a/Mathlib/Order/SetNotation.lean +++ b/Mathlib/Order/SetNotation.lean @@ -88,7 +88,7 @@ section delaborators open Lean Lean.PrettyPrinter.Delaborator /-- Delaborator for indexed supremum. -/ -@[delab app.iSup] +@[app_delab iSup] def iSup_delab : Delab := whenPPOption Lean.getPPNotation <| withOverApp 4 do let #[_, ι, _, f] := (← SubExpr.getExpr).getAppArgs | failure unless f.isLambda do failure @@ -116,7 +116,7 @@ def iSup_delab : Delab := whenPPOption Lean.getPPNotation <| withOverApp 4 do return stx /-- Delaborator for indexed infimum. -/ -@[delab app.iInf] +@[app_delab iInf] def iInf_delab : Delab := whenPPOption Lean.getPPNotation <| withOverApp 4 do let #[_, ι, _, f] := (← SubExpr.getExpr).getAppArgs | failure unless f.isLambda do failure @@ -193,7 +193,7 @@ section delaborators open Lean Lean.PrettyPrinter.Delaborator /-- Delaborator for indexed unions. -/ -@[delab app.Set.iUnion] +@[app_delab Set.iUnion] def iUnion_delab : Delab := whenPPOption Lean.getPPNotation do let #[_, ι, f] := (← SubExpr.getExpr).getAppArgs | failure unless f.isLambda do failure @@ -221,7 +221,7 @@ def iUnion_delab : Delab := whenPPOption Lean.getPPNotation do return stx /-- Delaborator for indexed intersections. -/ -@[delab app.Set.iInter] +@[app_delab Set.iInter] def sInter_delab : Delab := whenPPOption Lean.getPPNotation do let #[_, ι, f] := (← SubExpr.getExpr).getAppArgs | failure unless f.isLambda do failure diff --git a/Mathlib/RingTheory/Derivation/DifferentialRing.lean b/Mathlib/RingTheory/Derivation/DifferentialRing.lean index f26d5aab07492..a0d6e22ba7f17 100644 --- a/Mathlib/RingTheory/Derivation/DifferentialRing.lean +++ b/Mathlib/RingTheory/Derivation/DifferentialRing.lean @@ -27,7 +27,7 @@ open Lean PrettyPrinter Delaborator SubExpr in A delaborator for the x′ notation. This is required because it's not direct function application, so the default delaborator doesn't work. -/ -@[delab app.DFunLike.coe] +@[app_delab DFunLike.coe] def delabDeriv : Delab := do let e ← getExpr guard <| e.isAppOfArity' ``DFunLike.coe 6 diff --git a/Mathlib/Util/Delaborators.lean b/Mathlib/Util/Delaborators.lean index b45007dc4e2af..e091a084b2c56 100644 --- a/Mathlib/Util/Delaborators.lean +++ b/Mathlib/Util/Delaborators.lean @@ -97,7 +97,7 @@ open Lean Parser Term PrettyPrinter Delaborator /-- Delaborator for existential quantifier, including extended binders. -/ -- TODO: reduce the duplication in this code -@[delab app.Exists] +@[app_delab Exists] def exists_delab : Delab := whenPPOption Lean.getPPNotation do let #[ι, f] := (← SubExpr.getExpr).getAppArgs | failure unless f.isLambda do failure @@ -159,7 +159,7 @@ end existential open Lean Lean.PrettyPrinter.Delaborator /-- Delaborator for `∉`. -/ -@[delab app.Not] def delab_not_in := whenPPOption Lean.getPPNotation do +@[app_delab Not] def delab_not_in := whenPPOption Lean.getPPNotation do let #[f] := (← SubExpr.getExpr).getAppArgs | failure guard <| f.isAppOfArity ``Membership.mem 5 let stx₁ ← SubExpr.withAppArg <| SubExpr.withNaryArg 3 delab From e8a30842b99b9828a09461fffeaa56cd74287dac Mon Sep 17 00:00:00 2001 From: Artie Khovanov Date: Sat, 28 Dec 2024 03:20:01 +0000 Subject: [PATCH 019/189] chore(Algebra): rename theorems for consistency (#20271) Taking chore material out of #16094 to make the diff cleaner. Moves: isSquare_one -> IsSquare.one isSquare_sq -> IsSquare.sq even_two_nsmul -> Even.two_nsmul isSumSq_sum_mul_self -> IsSumSq.mul_self --- Mathlib/Algebra/Group/Even.lean | 9 +++++++-- Mathlib/Algebra/Ring/SumsOfSquares.lean | 4 +++- Mathlib/Data/Rat/Lemmas.lean | 4 ++-- .../RootSystem/Finite/CanonicalBilinear.lean | 2 +- .../NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean | 4 ++-- 5 files changed, 15 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/Group/Even.lean b/Mathlib/Algebra/Group/Even.lean index fa011c5c9b798..32a6d95f0fadf 100644 --- a/Mathlib/Algebra/Group/Even.lean +++ b/Mathlib/Algebra/Group/Even.lean @@ -88,7 +88,9 @@ instance Multiplicative.instDecidablePredIsSquare [DecidablePred (Even : α → end Add @[to_additive (attr := simp)] -lemma isSquare_one [MulOneClass α] : IsSquare (1 : α) := ⟨1, (mul_one _).symm⟩ +lemma IsSquare.one [MulOneClass α] : IsSquare (1 : α) := ⟨1, (mul_one _).symm⟩ + +@[to_additive, deprecated (since := "2024-12-27")] alias isSquare_one := IsSquare.one @[to_additive] lemma IsSquare.map [MulOneClass α] [MulOneClass β] [FunLike F α β] [MonoidHomClass F α β] @@ -114,7 +116,10 @@ attribute [to_additive Even.exists_two_nsmul @[to_additive Even.nsmul'] lemma Even.isSquare_pow : Even n → ∀ a : α, IsSquare (a ^ n) := by rintro ⟨n, rfl⟩ a; exact ⟨a ^ n, pow_add _ _ _⟩ -@[to_additive even_two_nsmul] lemma IsSquare_sq (a : α) : IsSquare (a ^ 2) := ⟨a, pow_two _⟩ +@[to_additive Even.two_nsmul] lemma IsSquare.sq (a : α) : IsSquare (a ^ 2) := ⟨a, pow_two _⟩ + +@[deprecated (since := "2024-12-27")] alias IsSquare_sq := IsSquare.sq +@[deprecated (since := "2024-12-27")] alias even_two_nsmul := Even.two_nsmul end Monoid diff --git a/Mathlib/Algebra/Ring/SumsOfSquares.lean b/Mathlib/Algebra/Ring/SumsOfSquares.lean index 6abec158dd334..83f2856af82f4 100644 --- a/Mathlib/Algebra/Ring/SumsOfSquares.lean +++ b/Mathlib/Algebra/Ring/SumsOfSquares.lean @@ -61,7 +61,7 @@ theorem IsSumSq.add [AddMonoid R] {S1 S2 : R} (p1 : IsSumSq S1) @[deprecated (since := "2024-08-09")] alias isSumSq.add := IsSumSq.add /-- A finite sum of squares is a sum of squares. -/ -theorem isSumSq_sum_mul_self {ι : Type*} [AddCommMonoid R] (s : Finset ι) (f : ι → R) : +theorem IsSumSq.sum_mul_self {ι : Type*} [AddCommMonoid R] (s : Finset ι) (f : ι → R) : IsSumSq (∑ i ∈ s, f i * f i) := by induction s using Finset.cons_induction with | empty => @@ -69,6 +69,8 @@ theorem isSumSq_sum_mul_self {ι : Type*} [AddCommMonoid R] (s : Finset ι) (f : | cons i s his h => exact (Finset.sum_cons (β := R) his) ▸ IsSumSq.sq_add (f i) (∑ i ∈ s, f i * f i) h +@[deprecated (since := "2024-12-27")] alias isSumSq_sum_mul_self := IsSumSq.sum_mul_self + variable (R) in /-- In an additive monoid with multiplication `R`, the type `sumSqIn R` is the submonoid of sums of diff --git a/Mathlib/Data/Rat/Lemmas.lean b/Mathlib/Data/Rat/Lemmas.lean index 06c65a1702cf3..739a18a81c1b1 100644 --- a/Mathlib/Data/Rat/Lemmas.lean +++ b/Mathlib/Data/Rat/Lemmas.lean @@ -119,11 +119,11 @@ theorem isSquare_iff {q : ℚ} : IsSquare q ↔ IsSquare q.num ∧ IsSquare q.de @[norm_cast, simp] theorem isSquare_natCast_iff {n : ℕ} : IsSquare (n : ℚ) ↔ IsSquare n := by - simp_rw [isSquare_iff, num_natCast, den_natCast, isSquare_one, and_true, Int.isSquare_natCast_iff] + simp_rw [isSquare_iff, num_natCast, den_natCast, IsSquare.one, and_true, Int.isSquare_natCast_iff] @[norm_cast, simp] theorem isSquare_intCast_iff {z : ℤ} : IsSquare (z : ℚ) ↔ IsSquare z := by - simp_rw [isSquare_iff, intCast_num, intCast_den, isSquare_one, and_true] + simp_rw [isSquare_iff, intCast_num, intCast_den, IsSquare.one, and_true] -- See note [no_index around OfNat.ofNat] @[simp] diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean index d291b7fbb148d..39b0b06f3214c 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean @@ -189,7 +189,7 @@ lemma corootForm_self_smul_root (i : ι) : lemma rootForm_self_sum_of_squares (x : M) : IsSumSq (P.RootForm x x) := - P.rootForm_apply_apply x x ▸ isSumSq_sum_mul_self Finset.univ _ + P.rootForm_apply_apply x x ▸ IsSumSq.sum_mul_self Finset.univ _ lemma rootForm_root_self (j : ι) : P.RootForm (P.root j) (P.root j) = ∑ (i : ι), (P.pairing j i) * (P.pairing j i) := by diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean index c99421dacd8f4..030cc20bab5bb 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean @@ -73,7 +73,7 @@ theorem quadraticCharFun_zero : quadraticCharFun F 0 = 0 := by @[simp] theorem quadraticCharFun_one : quadraticCharFun F 1 = 1 := by - simp only [quadraticCharFun, one_ne_zero, isSquare_one, if_true, if_false] + simp only [quadraticCharFun, one_ne_zero, IsSquare.one, if_true, if_false] /-- If `ringChar F = 2`, then `quadraticCharFun F` takes the value `1` on nonzero elements. -/ theorem quadraticCharFun_eq_one_of_char_two (hF : ringChar F = 2) {a : F} (ha : a ≠ 0) : @@ -140,7 +140,7 @@ theorem quadraticChar_one_iff_isSquare {a : F} (ha : a ≠ 0) : /-- The quadratic character takes the value `1` on nonzero squares. -/ theorem quadraticChar_sq_one' {a : F} (ha : a ≠ 0) : quadraticChar F (a ^ 2) = 1 := by - simp only [quadraticChar_apply, quadraticCharFun, sq_eq_zero_iff, ha, IsSquare_sq, if_true, + simp only [quadraticChar_apply, quadraticCharFun, sq_eq_zero_iff, ha, IsSquare.sq, if_true, if_false] /-- The square of the quadratic character on nonzero arguments is `1`. -/ From d09b8698f27b39bca5a463e40bb5b9d15200f0da Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Sat, 28 Dec 2024 10:33:18 +0000 Subject: [PATCH 020/189] chore: rename iteratedDeriv_const_{smul,mul} (#20136) --- Mathlib/Analysis/Calculus/Deriv/Basic.lean | 5 ++ Mathlib/Analysis/Calculus/FDeriv/Basic.lean | 7 ++ .../Calculus/IteratedDeriv/Lemmas.lean | 66 +++++++++++-------- .../Analysis/SpecialFunctions/ExpDeriv.lean | 4 +- 4 files changed, 54 insertions(+), 28 deletions(-) diff --git a/Mathlib/Analysis/Calculus/Deriv/Basic.lean b/Mathlib/Analysis/Calculus/Deriv/Basic.lean index de5b93a64a391..48874c1c046ed 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Basic.lean @@ -615,6 +615,7 @@ theorem deriv_id : deriv id x = 1 := theorem deriv_id' : deriv (@id 𝕜) = fun _ => 1 := funext deriv_id +/-- Variant with `fun x => x` rather than `id` -/ @[simp] theorem deriv_id'' : (deriv fun x : 𝕜 => x) = fun _ => 1 := deriv_id' @@ -622,6 +623,10 @@ theorem deriv_id'' : (deriv fun x : 𝕜 => x) = fun _ => 1 := theorem derivWithin_id (hxs : UniqueDiffWithinAt 𝕜 s x) : derivWithin id s x = 1 := (hasDerivWithinAt_id x s).derivWithin hxs +/-- Variant with `fun x => x` rather than `id` -/ +theorem derivWithin_id' (hxs : UniqueDiffWithinAt 𝕜 s x) : derivWithin (fun x => x) s x = 1 := + derivWithin_id x s hxs + end id section Const diff --git a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean index 58c5ca11932a3..101108c7507b4 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean @@ -1022,6 +1022,7 @@ theorem hasFDerivAt_id (x : E) : HasFDerivAt id (id 𝕜 E) x := theorem differentiableAt_id : DifferentiableAt 𝕜 id x := (hasFDerivAt_id x).differentiableAt +/-- Variant with `fun x => x` rather than `id` -/ @[simp] theorem differentiableAt_id' : DifferentiableAt 𝕜 (fun x => x) x := (hasFDerivAt_id x).differentiableAt @@ -1030,9 +1031,15 @@ theorem differentiableAt_id' : DifferentiableAt 𝕜 (fun x => x) x := theorem differentiableWithinAt_id : DifferentiableWithinAt 𝕜 id s x := differentiableAt_id.differentiableWithinAt +/-- Variant with `fun x => x` rather than `id` -/ +@[fun_prop] +theorem differentiableWithinAt_id' : DifferentiableWithinAt 𝕜 (fun x => x) s x := + differentiableWithinAt_id + @[simp, fun_prop] theorem differentiable_id : Differentiable 𝕜 (id : E → E) := fun _ => differentiableAt_id +/-- Variant with `fun x => x` rather than `id` -/ @[simp] theorem differentiable_id' : Differentiable 𝕜 fun x : E => x := fun _ => differentiableAt_id diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean index 1c58f1a462e70..c10db8c308277 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean @@ -65,16 +65,12 @@ theorem iteratedDerivWithin_const_sub (hn : 0 < n) (c : F) : @[deprecated (since := "2024-12-10")] alias iteratedDerivWithin_const_neg := iteratedDerivWithin_const_sub -/-- Note: this is unrelated to `iteratedDeriv_const_smul`, where the scalar multiplication works on -the domain. -/ theorem iteratedDerivWithin_const_smul (c : R) (hf : ContDiffOn 𝕜 n f s) : iteratedDerivWithin n (c • f) s x = c • iteratedDerivWithin n f s x := by simp_rw [iteratedDerivWithin] rw [iteratedFDerivWithin_const_smul_apply hf h hx] simp only [ContinuousMultilinearMap.smul_apply] -/-- Note: this is unrelated to `iteratedDeriv_const_mul`, where the multiplication works on the -domain. -/ theorem iteratedDerivWithin_const_mul (c : 𝕜) {f : 𝕜 → 𝕜} (hf : ContDiffOn 𝕜 n f s) : iteratedDerivWithin n (fun z => c * f z) s x = c * iteratedDerivWithin n f s x := by simpa using iteratedDerivWithin_const_smul (F := 𝕜) hx h c hf @@ -96,6 +92,32 @@ 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] +theorem iteratedDerivWithin_comp_const_smul (hf : ContDiffOn 𝕜 n f s) (c : 𝕜) + (hs : Set.MapsTo (c * ·) s s) : + iteratedDerivWithin n (fun x => f (c * x)) s x = c ^ n • iteratedDerivWithin n f s (c * x) := by + induction n generalizing x with + | zero => simp + | succ n ih => + have hcx : c * x ∈ s := hs hx + have h₀ : s.EqOn + (iteratedDerivWithin n (fun x ↦ f (c * x)) s) + (fun x => c ^ n • iteratedDerivWithin n f s (c * x)) := + fun x hx => ih hx hf.of_succ + have h₁ : DifferentiableWithinAt 𝕜 (iteratedDerivWithin n f s) s (c * x) := + hf.differentiableOn_iteratedDerivWithin (Nat.cast_lt.mpr n.lt_succ_self) h _ hcx + have h₂ : DifferentiableWithinAt 𝕜 (fun x => iteratedDerivWithin n f s (c * x)) s x := by + rw [← Function.comp_def] + apply DifferentiableWithinAt.comp + · exact hf.differentiableOn_iteratedDerivWithin (Nat.cast_lt.mpr n.lt_succ_self) h _ hcx + · exact differentiableWithinAt_id'.const_mul _ + · exact hs + rw [iteratedDerivWithin_succ (h _ hx), derivWithin_congr h₀ (ih hx hf.of_succ), + derivWithin_const_smul (h _ hx) (c ^ n) h₂, iteratedDerivWithin_succ (h _ hcx), + ← Function.comp_def, + derivWithin.scomp x h₁ (differentiableWithinAt_id'.const_mul _) hs (h _ hx), + derivWithin_const_mul (h _ hx) _ differentiableWithinAt_id', derivWithin_id' _ _ (h _ hx), + smul_smul, mul_one, pow_succ] + end lemma iteratedDeriv_add (hf : ContDiff 𝕜 n f) (hg : ContDiff 𝕜 n g) : @@ -125,30 +147,22 @@ lemma iteratedDeriv_sub (hf : ContDiff 𝕜 n f) (hg : ContDiff 𝕜 n g) : iteratedDerivWithin_sub (Set.mem_univ _) uniqueDiffOn_univ (contDiffOn_univ.mpr hf) (contDiffOn_univ.mpr hg) -/-- Note: this is unrelated to `iteratedDerivWithin_const_smul`, where the scalar multiplication -works on the codomain. -/ -theorem iteratedDeriv_const_smul {n : ℕ} {f : 𝕜 → F} (h : ContDiff 𝕜 n f) (c : 𝕜) : +theorem iteratedDeriv_comp_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 - | zero => simp - | succ n ih => - funext x - have h₀ : DifferentiableAt 𝕜 (iteratedDeriv n f) (c * x) := - h.differentiable_iteratedDeriv n (Nat.cast_lt.mpr n.lt_succ_self) |>.differentiableAt - have h₁ : DifferentiableAt 𝕜 (fun x => iteratedDeriv n f (c * x)) x := by - rw [← Function.comp_def] - apply DifferentiableAt.comp - · exact h.differentiable_iteratedDeriv n (Nat.cast_lt.mpr n.lt_succ_self) |>.differentiableAt - · exact differentiableAt_id'.const_mul _ - rw [iteratedDeriv_succ, ih h.of_succ, deriv_const_smul _ h₁, iteratedDeriv_succ, - ← Function.comp_def, deriv.scomp x h₀ (differentiableAt_id'.const_mul _), - deriv_const_mul _ differentiableAt_id', deriv_id'', smul_smul, mul_one, pow_succ] - -/-- Note: this is unrelated to `iteratedDerivWithin_const_mul`, where the multiplication works on -the codomain. -/ -theorem iteratedDeriv_const_mul {n : ℕ} {f : 𝕜 → 𝕜} (h : ContDiff 𝕜 n f) (c : 𝕜) : + funext x + simpa only [iteratedDerivWithin_univ] using + iteratedDerivWithin_comp_const_smul (Set.mem_univ x) uniqueDiffOn_univ (contDiffOn_univ.mpr h) + c (Set.mapsTo_univ _ _) + +@[deprecated (since := "2024-12-20")] +alias iteratedDeriv_const_smul := iteratedDeriv_comp_const_smul + +theorem iteratedDeriv_comp_const_mul {n : ℕ} {f : 𝕜 → 𝕜} (h : ContDiff 𝕜 n f) (c : 𝕜) : iteratedDeriv n (fun x => f (c * x)) = fun x => c ^ n * iteratedDeriv n f (c * x) := by - simpa only [smul_eq_mul] using iteratedDeriv_const_smul h c + simpa only [smul_eq_mul] using iteratedDeriv_comp_const_smul h c + +@[deprecated (since := "2024-12-20")] +alias iteratedDeriv_const_mul := iteratedDeriv_comp_const_mul lemma iteratedDeriv_comp_neg (n : ℕ) (f : 𝕜 → F) (a : 𝕜) : iteratedDeriv n (fun x ↦ f (-x)) a = (-1 : 𝕜) ^ n • iteratedDeriv n f (-a) := by diff --git a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean index 6a54e8ef68876..accf98401aed5 100644 --- a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean @@ -187,7 +187,7 @@ open Complex in @[simp] theorem iteratedDeriv_cexp_const_mul (n : ℕ) (c : ℂ) : (iteratedDeriv n fun s : ℂ => exp (c * s)) = fun s => c ^ n * exp (c * s) := by - rw [iteratedDeriv_const_mul contDiff_exp, iteratedDeriv_eq_iterate, iter_deriv_exp] + rw [iteratedDeriv_comp_const_mul contDiff_exp, iteratedDeriv_eq_iterate, iter_deriv_exp] /-! ## `Real.exp` -/ @@ -351,4 +351,4 @@ open Real in @[simp] theorem iteratedDeriv_exp_const_mul (n : ℕ) (c : ℝ) : (iteratedDeriv n fun s => exp (c * s)) = fun s => c ^ n * exp (c * s) := by - rw [iteratedDeriv_const_mul contDiff_exp, iteratedDeriv_eq_iterate, iter_deriv_exp] + rw [iteratedDeriv_comp_const_mul contDiff_exp, iteratedDeriv_eq_iterate, iter_deriv_exp] From 161c83c3e874142a5a1658146ac3d26e43a91dc6 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sat, 28 Dec 2024 11:04:06 +0000 Subject: [PATCH 021/189] feat(RingTheory): Jacobi-Zariski sequence (#19067) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + .../Module/Presentation/Differentials.lean | 4 +- Mathlib/Data/Finsupp/Basic.lean | 23 +- Mathlib/RingTheory/Generators.lean | 93 +++- .../RingTheory/Kaehler/CotangentComplex.lean | 8 +- Mathlib/RingTheory/Kaehler/JacobiZariski.lean | 492 ++++++++++++++++++ 6 files changed, 604 insertions(+), 17 deletions(-) create mode 100644 Mathlib/RingTheory/Kaehler/JacobiZariski.lean diff --git a/Mathlib.lean b/Mathlib.lean index 19697cf5f9461..c0bbcc04a8b3b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4456,6 +4456,7 @@ import Mathlib.RingTheory.Jacobson.Polynomial import Mathlib.RingTheory.Jacobson.Ring import Mathlib.RingTheory.Kaehler.Basic import Mathlib.RingTheory.Kaehler.CotangentComplex +import Mathlib.RingTheory.Kaehler.JacobiZariski import Mathlib.RingTheory.Kaehler.Polynomial import Mathlib.RingTheory.Kaehler.TensorProduct import Mathlib.RingTheory.KrullDimension.Basic diff --git a/Mathlib/Algebra/Module/Presentation/Differentials.lean b/Mathlib/Algebra/Module/Presentation/Differentials.lean index 5ec008c5b1ed6..27121d35ea0d8 100644 --- a/Mathlib/Algebra/Module/Presentation/Differentials.lean +++ b/Mathlib/Algebra/Module/Presentation/Differentials.lean @@ -68,8 +68,8 @@ lemma comm₂₃' : pres.toExtension.toKaehler.comp pres.cotangentSpaceBasis.rep dsimp rw [Basis.repr_symm_apply, Finsupp.linearCombination_single, Finsupp.linearCombination_single, one_smul, one_smul, - Generators.cotangentSpaceBasis_apply, mapBaseChange_tmul, one_smul] - simp + Generators.cotangentSpaceBasis_apply] + simp [Generators.toExtension] /-- The canonical map `(pres.rels →₀ S) →ₗ[S] pres.toExtension.Cotangent`. -/ noncomputable def hom₁ : (pres.rels →₀ S) →ₗ[S] pres.toExtension.Cotangent := diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index 7c786c4ec0d7e..8f9546c197a70 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -1124,18 +1124,11 @@ end CurryUncurry section Sum /-- `Finsupp.sumElim f g` maps `inl x` to `f x` and `inr y` to `g y`. -/ -def sumElim {α β γ : Type*} [Zero γ] (f : α →₀ γ) (g : β →₀ γ) : α ⊕ β →₀ γ := - onFinset - (by - haveI := Classical.decEq α - haveI := Classical.decEq β - exact f.support.map ⟨_, Sum.inl_injective⟩ ∪ g.support.map ⟨_, Sum.inr_injective⟩) - (Sum.elim f g) fun ab h => by - cases' ab with a b <;> - letI := Classical.decEq α <;> letI := Classical.decEq β <;> - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/10754): had to add these `DecidableEq` instances - simp only [Sum.elim_inl, Sum.elim_inr] at h <;> - simpa +@[simps support] +def sumElim {α β γ : Type*} [Zero γ] (f : α →₀ γ) (g : β →₀ γ) : α ⊕ β →₀ γ where + support := f.support.disjSum g.support + toFun := Sum.elim f g + mem_support_toFun := by simp @[simp, norm_cast] theorem coe_sumElim {α β γ : Type*} [Zero γ] (f : α →₀ γ) (g : β →₀ γ) : @@ -1154,6 +1147,12 @@ theorem sumElim_inr {α β γ : Type*} [Zero γ] (f : α →₀ γ) (g : β → sumElim f g (Sum.inr x) = g x := rfl +@[to_additive] +lemma prod_sumElim {ι₁ ι₂ α M : Type*} [Zero α] [CommMonoid M] + (f₁ : ι₁ →₀ α) (f₂ : ι₂ →₀ α) (g : ι₁ ⊕ ι₂ → α → M) : + (f₁.sumElim f₂).prod g = f₁.prod (g ∘ Sum.inl) * f₂.prod (g ∘ Sum.inr) := by + simp [Finsupp.prod, Finset.prod_disj_sum] + /-- The equivalence between `(α ⊕ β) →₀ γ` and `(α →₀ γ) × (β →₀ γ)`. This is the `Finsupp` version of `Equiv.sum_arrow_equiv_prod_arrow`. -/ diff --git a/Mathlib/RingTheory/Generators.lean b/Mathlib/RingTheory/Generators.lean index aa91a2d21ba7c..d63cf6d23df0c 100644 --- a/Mathlib/RingTheory/Generators.lean +++ b/Mathlib/RingTheory/Generators.lean @@ -367,6 +367,9 @@ def toComp (Q : Generators S T) (P : Generators R S) : Hom P (Q.comp P) where val i := X (.inr i) aeval_val i := by simp +lemma toComp_toAlgHom (Q : Generators S T) (P : Generators R S) : + (Q.toComp P).toAlgHom = rename Sum.inr := rfl + /-- Given families of generators `X ⊆ T` over `S` and `Y ⊆ S` over `R`, there is a map of generators `R[X, Y] → S[X]`. -/ @[simps] @@ -375,6 +378,24 @@ def ofComp (Q : Generators S T) (P : Generators R S) : Hom (Q.comp P) Q where val i := i.elim X (C ∘ P.val) aeval_val i := by cases i <;> simp +lemma ofComp_toAlgHom_monomial_sumElim (Q : Generators S T) (P : Generators R S) (v₁ v₂ a) : + (Q.ofComp P).toAlgHom (monomial (Finsupp.sumElim v₁ v₂) a) = + monomial v₁ (aeval P.val (monomial v₂ a)) := by + erw [Hom.toAlgHom_monomial] + rw [monomial_eq] + simp only [MvPolynomial.algebraMap_apply, ofComp_val, aeval_monomial] + rw [Finsupp.prod_sumElim] + simp only [Function.comp_def, Sum.elim_inl, Sum.elim_inr, ← map_pow, ← map_finsupp_prod, + C_mul, algebra.smul_def, MvPolynomial.algebraMap_apply, mul_assoc] + nth_rw 2 [mul_comm] + +lemma toComp_toAlgHom_monomial (Q : Generators S T) (P : Generators R S) (j a) : + (Q.toComp P).toAlgHom (monomial j a) = + monomial (Finsupp.sumElim 0 j) a := by + convert rename_monomial _ _ _ + ext f (i₁ | i₂) <;> + simp [Finsupp.mapDomain_notin_range, Finsupp.mapDomain_apply Sum.inr_injective] + /-- Given families of generators `X ⊆ T`, there is a map `R[X] → S[X]`. -/ @[simps] noncomputable @@ -403,12 +424,80 @@ lemma Hom.toExtensionHom_comp [Algebra R S'] [IsScalarTower R S S'] (f : P'.Hom P'') (g : P.Hom P') : toExtensionHom (f.comp g) = f.toExtensionHom.comp g.toExtensionHom := by ext; simp -end Hom - /-- The kernel of a presentation. -/ noncomputable abbrev ker : Ideal P.Ring := P.toExtension.ker lemma ker_eq_ker_aeval_val : P.ker = RingHom.ker (aeval P.val) := rfl +lemma map_toComp_ker (Q : Generators S T) (P : Generators R S) : + P.ker.map (Q.toComp P).toAlgHom.toRingHom = RingHom.ker (Q.ofComp P).toAlgHom := by + letI : DecidableEq (Q.vars →₀ ℕ) := Classical.decEq _ + apply le_antisymm + · rw [Ideal.map_le_iff_le_comap] + rintro x (hx : algebraMap P.Ring S x = 0) + have : (Q.ofComp P).toAlgHom.comp (Q.toComp P).toAlgHom = IsScalarTower.toAlgHom R _ _ := by + ext1; simp + simp only [comp_vars, AlgHom.toRingHom_eq_coe, Ideal.mem_comap, RingHom.coe_coe, + RingHom.mem_ker, ← AlgHom.comp_apply, this, IsScalarTower.toAlgHom_apply] + rw [IsScalarTower.algebraMap_apply P.Ring S, hx, map_zero] + · rintro x (h₂ : (Q.ofComp P).toAlgHom x = 0) + let e : ((Q.comp P).vars →₀ ℕ) ≃+ (Q.vars →₀ ℕ) × (P.vars →₀ ℕ) := + Finsupp.sumFinsuppAddEquivProdFinsupp + suffices ∑ v ∈ (support x).map e, (monomial (e.symm v)) (coeff (e.symm v) x) ∈ + Ideal.map (Q.toComp P).toAlgHom.toRingHom P.ker by + simpa only [AlgHom.toRingHom_eq_coe, Finset.sum_map, Equiv.coe_toEmbedding, + EquivLike.coe_coe, AddEquiv.symm_apply_apply, support_sum_monomial_coeff] using this + rw [← Finset.sum_fiberwise_of_maps_to (fun i ↦ Finset.mem_image_of_mem Prod.fst)] + refine sum_mem fun i hi ↦ ?_ + convert_to monomial (e.symm (i, 0)) 1 * (Q.toComp P).toAlgHom.toRingHom + (∑ j ∈ ((support x).map e.toEmbedding).filter (fun x ↦ x.1 = i), + monomial j.2 (coeff (e.symm j) x)) ∈ _ + · rw [map_sum, Finset.mul_sum] + refine Finset.sum_congr rfl fun j hj ↦ ?_ + obtain rfl := (Finset.mem_filter.mp hj).2 + obtain ⟨i, j⟩ := j + clear hj hi + have : (Q.toComp P).toAlgHom (monomial j (coeff (e.symm (i, j)) x)) = + monomial (e.symm (0, j)) (coeff (e.symm (i, j)) x) := + toComp_toAlgHom_monomial .. + simp only [AlgHom.toRingHom_eq_coe, monomial_zero', RingHom.coe_coe, algHom_C, + MvPolynomial.algebraMap_eq, this] + rw [monomial_mul, ← map_add, Prod.mk_add_mk, add_zero, zero_add, one_mul] + · apply Ideal.mul_mem_left + refine Ideal.mem_map_of_mem _ ?_ + simp only [ker_eq_ker_aeval_val, AddEquiv.toEquiv_eq_coe, RingHom.mem_ker, map_sum] + rw [← coeff_zero i, ← h₂] + clear h₂ hi + have (x : (Q.comp P).Ring) : (Function.support fun a ↦ if a.1 = i then aeval P.val + (monomial a.2 (coeff (e.symm a) x)) else 0) ⊆ ((support x).map e).toSet := by + rw [← Set.compl_subset_compl] + intro j + obtain ⟨j, rfl⟩ := e.surjective j + simp_all + rw [Finset.sum_filter, ← finsum_eq_sum_of_support_subset _ (this x)] + induction x using MvPolynomial.induction_on' with + | h1 v a => + rw [finsum_eq_sum_of_support_subset _ (this _), ← Finset.sum_filter] + obtain ⟨v, rfl⟩ := e.symm.surjective v + erw [ofComp_toAlgHom_monomial_sumElim] + classical + simp only [comp_vars, coeff_monomial, ← e.injective.eq_iff, + map_zero, AddEquiv.apply_symm_apply, apply_ite] + rw [← apply_ite, Finset.sum_ite_eq] + simp only [Finset.mem_filter, Finset.mem_map_equiv, AddEquiv.coe_toEquiv_symm, comp_vars, + mem_support_iff, coeff_monomial, ↓reduceIte, ne_eq, ite_and, ite_not] + split + · simp only [zero_smul, coeff_zero, *, map_zero, ite_self] + · congr + | h2 p q hp hq => + simp only [coeff_add, map_add, ite_add_zero] + rw [finsum_add_distrib, hp, hq] + · refine (((support p).map e).finite_toSet.subset ?_) + convert this p + · refine (((support q).map e).finite_toSet.subset ?_) + convert this q + +end Hom + end Algebra.Generators diff --git a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean index 91c2fb12f48ba..79c48f487358b 100644 --- a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean +++ b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean @@ -377,7 +377,7 @@ lemma cotangentSpaceBasis_repr_one_tmul (x i) : simp lemma cotangentSpaceBasis_apply (i) : - P.cotangentSpaceBasis i = 1 ⊗ₜ .D _ _ (.X i) := by + P.cotangentSpaceBasis i = ((1 : S) ⊗ₜ[P.Ring] D R P.Ring (.X i) : _) := by simp [cotangentSpaceBasis, toExtension] universe w' u' v' @@ -411,6 +411,12 @@ lemma repr_CotangentSpaceMap (f : Hom P P') (i j) : rw [CotangentSpace.map_tmul, map_one] erw [cotangentSpaceBasis_repr_one_tmul, Hom.toAlgHom_X] +@[simp] +lemma toKaehler_cotangentSpaceBasis (i) : + P.toExtension.toKaehler (P.cotangentSpaceBasis i) = D R S (P.val i) := by + rw [cotangentSpaceBasis_apply] + exact (KaehlerDifferential.mapBaseChange_tmul ..).trans (by simp) + end Generators variable {P : Generators R S} diff --git a/Mathlib/RingTheory/Kaehler/JacobiZariski.lean b/Mathlib/RingTheory/Kaehler/JacobiZariski.lean new file mode 100644 index 0000000000000..c789134523f00 --- /dev/null +++ b/Mathlib/RingTheory/Kaehler/JacobiZariski.lean @@ -0,0 +1,492 @@ +/- +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.CotangentComplex +import Mathlib.RingTheory.Generators +import Mathlib.Algebra.Module.SnakeLemma + +/-! + +# The Jacobi-Zariski exact sequence + +Given `R → S → T`, the Jacobi-Zariski exact sequence is +``` +H¹(L_{T/R}) → H¹(L_{T/S}) → T ⊗[S] Ω[S/R] → Ω[T/R] → Ω[T/S] → 0 +``` +The maps are +- `Algebra.H1Cotangent.map` +- `Algebra.H1Cotangent.δ` +- `KaehlerDifferential.mapBaseChange` +- `KaehlerDifferential.map` +and the exactness lemmas are +- `Algebra.H1Cotangent.exact_map_δ` +- `Algebra.H1Cotangent.exact_δ_mapBaseChange` +- `KaehlerDifferential.exact_mapBaseChange_map` +- `KaehlerDifferential.map_surjective` +-/ + +open KaehlerDifferential TensorProduct MvPolynomial + +namespace Algebra + +universe u₁ u₂ u₃ u₄ w' w u v uT + +variable {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Generators.{w} R S} +variable {T : Type uT} [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] +variable (Q : Generators.{w} S T) (P : Generators.{w'} R S) + +attribute [local instance] SMulCommClass.of_commMonoid + +namespace Generators + +/-- +Given `R[X] → S` and `S[Y] → T`, this is the lift of an element in `ker(S[Y] → T)` +to `ker(R[X][Y] → S[Y] → T)` constructed from `P.σ`. +-/ +noncomputable +def kerCompPreimage (x : Q.ker) : + (Q.comp P).ker := by + refine ⟨x.1.sum fun n r ↦ ?_, ?_⟩ + · -- The use of `refine` is intentional to control the elaboration order + -- so that the term has type `(Q.comp P).Ring` and not `MvPolynomial (Q.vars ⊕ P.vars) R` + refine rename ?_ (P.σ r) * monomial ?_ 1 + exacts [Sum.inr, n.mapDomain Sum.inl] + · simp only [ker_eq_ker_aeval_val, RingHom.mem_ker] + conv_rhs => rw [← show aeval Q.val x.1 = 0 from x.2, ← x.1.support_sum_monomial_coeff] + simp only [Finsupp.sum, map_sum, map_mul, aeval_rename, Function.comp_def, comp_val, + Sum.elim_inr, aeval_monomial, map_one, Finsupp.prod_mapDomain_index_inj Sum.inl_injective, + Sum.elim_inl, one_mul] + congr! with v i + simp_rw [← IsScalarTower.toAlgHom_apply R, ← comp_aeval, AlgHom.comp_apply, P.aeval_val_σ] + rfl + +lemma ofComp_kerCompPreimage (x : Q.ker) : + (Q.ofComp P).toAlgHom (kerCompPreimage Q P x) = x := by + conv_rhs => rw [← x.1.support_sum_monomial_coeff] + rw [kerCompPreimage, map_finsupp_sum, Finsupp.sum] + refine Finset.sum_congr rfl fun j _ ↦ ?_ + simp only [AlgHom.toLinearMap_apply, _root_.map_mul, Hom.toAlgHom_monomial] + rw [one_smul, Finsupp.prod_mapDomain_index_inj Sum.inl_injective] + rw [rename, ← AlgHom.comp_apply, comp_aeval] + simp only [ofComp_val, Sum.elim_inr, Function.comp_apply, self_val, id_eq, + Sum.elim_inl, monomial_eq, Hom.toAlgHom_X] + congr 1 + rw [aeval_def, IsScalarTower.algebraMap_eq R S, ← MvPolynomial.algebraMap_eq, + ← coe_eval₂Hom, ← map_aeval, P.aeval_val_σ] + rfl + +lemma Cotangent.map_ofComp_ker : + Submodule.map (Q.ofComp P).toAlgHom.toLinearMap ((Q.comp P).ker.restrictScalars R) = + Q.ker.restrictScalars R := by + apply le_antisymm + · rintro _ ⟨x, hx, rfl⟩ + simp only [ker_eq_ker_aeval_val, Submodule.coe_restrictScalars, SetLike.mem_coe, + RingHom.mem_ker, AlgHom.toLinearMap_apply, Submodule.restrictScalars_mem] at hx ⊢ + rw [← hx, Hom.algebraMap_toAlgHom] + rfl + · intro x hx + exact ⟨_, (kerCompPreimage Q P ⟨x, hx⟩).2, ofComp_kerCompPreimage Q P ⟨x, hx⟩⟩ + +lemma Cotangent.surjective_map_ofComp : + Function.Surjective (Extension.Cotangent.map (Q.ofComp P).toExtensionHom) := by + intro x + obtain ⟨⟨x, hx⟩, rfl⟩ := Extension.Cotangent.mk_surjective x + have : x ∈ Q.ker.restrictScalars R := hx + rw [← map_ofComp_ker Q P] at this + obtain ⟨x, hx', rfl⟩ := this + exact ⟨.mk ⟨x, hx'⟩, Extension.Cotangent.map_mk _ _⟩ + +/-! +Given representations `0 → I → R[X] → S → 0` and `0 → K → S[Y] → T → 0`, +we may consider the induced representation `0 → J → R[X, Y] → T → 0`, and the sequence +`T ⊗[S] (I/I²) → J/J² → K/K²` is exact. +-/ +open Extension.Cotangent in +lemma Cotangent.exact : + Function.Exact + ((Extension.Cotangent.map (Q.toComp P).toExtensionHom).liftBaseChange T) + (Extension.Cotangent.map (Q.ofComp P).toExtensionHom) := by + apply LinearMap.exact_of_comp_of_mem_range + · rw [LinearMap.liftBaseChange_comp, ← Extension.Cotangent.map_comp, + EmbeddingLike.map_eq_zero_iff] + ext x + obtain ⟨⟨x, hx⟩, rfl⟩ := Extension.Cotangent.mk_surjective x + simp only [map_mk, Hom.toAlgHom_comp_apply, val_mk, LinearMap.zero_apply, val_zero] + convert Q.ker.toCotangent.map_zero + trans ((IsScalarTower.toAlgHom R _ _).comp (IsScalarTower.toAlgHom R P.Ring S)) x + · congr + refine MvPolynomial.algHom_ext fun i ↦ ?_ + show (Q.ofComp P).toAlgHom ((Q.toComp P).toAlgHom (X i)) = _ + simp + · simp [-self_vars, show aeval P.val x = 0 from hx] + · intro x hx + obtain ⟨⟨x : (Q.comp P).Ring, hx'⟩, rfl⟩ := Extension.Cotangent.mk_surjective x + replace hx : (Q.ofComp P).toAlgHom x ∈ Q.ker ^ 2 := by + simpa only [map_mk, val_mk, val_zero, Ideal.toCotangent_eq_zero] using congr(($hx).val) + rw [← Submodule.restrictScalars_mem R, pow_two, Submodule.restrictScalars_mul, + ← map_ofComp_ker (P := P), ← Submodule.map_mul, ← Submodule.restrictScalars_mul] at hx + obtain ⟨y, hy, e⟩ := hx + rw [AlgHom.toLinearMap_apply, eq_comm, ← sub_eq_zero, ← map_sub, ← RingHom.mem_ker, + ← map_toComp_ker] at e + rw [LinearMap.range_liftBaseChange] + let z : (Q.comp P).ker := ⟨x - y, Ideal.sub_mem _ hx' (Ideal.mul_le_left hy)⟩ + have hz : z.1 ∈ P.ker.map (Q.toComp P).toAlgHom.toRingHom := e + have : Extension.Cotangent.mk (P := (Q.comp P).toExtension) ⟨x, hx'⟩ = + Extension.Cotangent.mk z := by + ext; simpa only [comp_vars, val_mk, Ideal.toCotangent_eq, sub_sub_cancel, pow_two] + rw [this, ← Submodule.restrictScalars_mem (Q.comp P).Ring, ← Submodule.mem_comap, + ← Submodule.span_singleton_le_iff_mem, ← Submodule.map_le_map_iff_of_injective + (f := Submodule.subtype _) Subtype.val_injective, Submodule.map_subtype_span_singleton, + Submodule.span_singleton_le_iff_mem] + refine (show Ideal.map (Q.toComp P).toAlgHom.toRingHom P.ker ≤ _ from ?_) hz + rw [Ideal.map_le_iff_le_comap] + rintro w hw + simp only [AlgHom.toRingHom_eq_coe, Ideal.mem_comap, RingHom.coe_coe, + Submodule.mem_map, Submodule.mem_comap, Submodule.restrictScalars_mem, Submodule.coe_subtype, + Subtype.exists, exists_and_right, exists_eq_right, + toExtension_Ring, toExtension_commRing, toExtension_algebra₂] + refine ⟨?_, Submodule.subset_span ⟨Extension.Cotangent.mk ⟨w, hw⟩, ?_⟩⟩ + · simp only [ker_eq_ker_aeval_val, RingHom.mem_ker, Hom.algebraMap_toAlgHom] + rw [show aeval P.val w = 0 from hw, map_zero] + · rw [map_mk] + rfl + +/-- Given `R[X] → S` and `S[Y] → T`, the cotangent space of `R[X][Y] → T` is isomorphic +to the direct product of the cotangent space of `S[Y] → T` and `R[X] → S` (base changed to `T`). -/ +noncomputable +def CotangentSpace.compEquiv (Q : Generators.{w} S T) (P : Generators.{w'} R S) : + (Q.comp P).toExtension.CotangentSpace ≃ₗ[T] + Q.toExtension.CotangentSpace × (T ⊗[S] P.toExtension.CotangentSpace) := + (Q.comp P).cotangentSpaceBasis.repr.trans + (Q.cotangentSpaceBasis.prod (P.cotangentSpaceBasis.baseChange T)).repr.symm + +section instanceProblem + +-- Note: these instances are needed to prevent instance search timeouts. +attribute [local instance 999999] Zero.toOfNat0 SemilinearMapClass.distribMulActionSemiHomClass + SemilinearEquivClass.instSemilinearMapClass TensorProduct.addZeroClass AddZeroClass.toZero + +lemma CotangentSpace.compEquiv_symm_inr : + (compEquiv Q P).symm.toLinearMap ∘ₗ + LinearMap.inr T Q.toExtension.CotangentSpace (T ⊗[S] P.toExtension.CotangentSpace) = + (Extension.CotangentSpace.map (Q.toComp P).toExtensionHom).liftBaseChange T := by + classical + apply (P.cotangentSpaceBasis.baseChange T).ext + intro i + apply (Q.comp P).cotangentSpaceBasis.repr.injective + ext j + simp only [compEquiv, LinearEquiv.trans_symm, LinearEquiv.symm_symm, + Basis.baseChange_apply, LinearMap.coe_comp, LinearEquiv.coe_coe, LinearMap.coe_inr, + Function.comp_apply, LinearEquiv.trans_apply, Basis.repr_symm_apply, pderiv_X, toComp_val, + Basis.repr_linearCombination, LinearMap.liftBaseChange_tmul, one_smul, repr_CotangentSpaceMap] + obtain (j | j) := j <;> + simp only [comp_vars, Basis.prod_repr_inr, Basis.baseChange_repr_tmul, + Basis.repr_self, Basis.prod_repr_inl, map_zero, Finsupp.coe_zero, + Pi.zero_apply, ne_eq, not_false_eq_true, Pi.single_eq_of_ne, Pi.single_apply, + Finsupp.single_apply, ite_smul, one_smul, zero_smul, Sum.inr.injEq, + RingHom.map_ite_one_zero, reduceCtorEq, ↓reduceIte] + +lemma CotangentSpace.compEquiv_symm_zero (x) : + (compEquiv Q P).symm (0, x) = + (Extension.CotangentSpace.map (Q.toComp P).toExtensionHom).liftBaseChange T x := + DFunLike.congr_fun (compEquiv_symm_inr Q P) x + +lemma CotangentSpace.fst_compEquiv : + LinearMap.fst T Q.toExtension.CotangentSpace (T ⊗[S] P.toExtension.CotangentSpace) ∘ₗ + (compEquiv Q P).toLinearMap = Extension.CotangentSpace.map (Q.ofComp P).toExtensionHom := by + classical + apply (Q.comp P).cotangentSpaceBasis.ext + intro i + apply Q.cotangentSpaceBasis.repr.injective + ext j + simp only [compEquiv, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, ofComp_val, + LinearEquiv.trans_apply, Basis.repr_self, LinearMap.fst_apply, repr_CotangentSpaceMap] + obtain (i | i) := i <;> + simp only [comp_vars, Basis.repr_symm_apply, Finsupp.linearCombination_single, Basis.prod_apply, + LinearMap.coe_inl, LinearMap.coe_inr, Sum.elim_inl, Function.comp_apply, one_smul, + Basis.repr_self, Finsupp.single_apply, pderiv_X, Pi.single_apply, RingHom.map_ite_one_zero, + Sum.elim_inr, Function.comp_apply, Basis.baseChange_apply, one_smul, + map_zero, Finsupp.coe_zero, Pi.zero_apply, derivation_C] + +lemma CotangentSpace.fst_compEquiv_apply (x) : + (compEquiv Q P x).1 = Extension.CotangentSpace.map (Q.ofComp P).toExtensionHom x := + DFunLike.congr_fun (fst_compEquiv Q P) x + +lemma CotangentSpace.map_toComp_injective : + Function.Injective + ((Extension.CotangentSpace.map (Q.toComp P).toExtensionHom).liftBaseChange T) := by + rw [← compEquiv_symm_inr] + apply (compEquiv Q P).symm.injective.comp + exact Prod.mk.inj_left _ + +lemma CotangentSpace.map_ofComp_surjective : + Function.Surjective (Extension.CotangentSpace.map (Q.ofComp P).toExtensionHom) := by + rw [← fst_compEquiv] + exact (Prod.fst_surjective).comp (compEquiv Q P).surjective + +/-! +Given representations `R[X] → S` and `S[Y] → T`, the sequence +`T ⊗[S] (⨁ₓ S dx) → (⨁ₓ T dx) ⊕ (⨁ᵧ T dy) → ⨁ᵧ T dy` +is exact. +-/ +lemma CotangentSpace.exact : + Function.Exact ((Extension.CotangentSpace.map (Q.toComp P).toExtensionHom).liftBaseChange T) + (Extension.CotangentSpace.map (Q.ofComp P).toExtensionHom) := by + rw [← fst_compEquiv, ← compEquiv_symm_inr] + conv_rhs => rw [← LinearEquiv.symm_symm (compEquiv Q P)] + rw [LinearEquiv.conj_exact_iff_exact] + exact Function.Exact.inr_fst + +namespace H1Cotangent + +variable (R) in +/-- +Given `0 → I → S[Y] → T → 0`, this is an auxiliary map from `S[Y]` to `T ⊗[S] Ω[S⁄R]` whose +restriction to `ker(I/I² → ⊕ S dyᵢ)` is the connecting homomorphism in the Jacobi-Zariski sequence. +-/ +noncomputable +def δAux : + Q.Ring →ₗ[R] T ⊗[S] Ω[S⁄R] := + Finsupp.lsum R (R := R) fun f ↦ + (TensorProduct.mk S T _ (f.prod (Q.val · ^ ·))).restrictScalars R ∘ₗ (D R S).toLinearMap + +lemma δAux_monomial (n r) : + δAux R Q (monomial n r) = (n.prod (Q.val · ^ ·)) ⊗ₜ D R S r := + Finsupp.lsum_single _ _ _ _ + +@[simp] +lemma δAux_X (i) : + δAux R Q (X i) = 0 := by + rw [X, δAux_monomial] + simp only [Derivation.map_one_eq_zero, tmul_zero] + +lemma δAux_mul (x y) : + δAux R Q (x * y) = x • (δAux R Q y) + y • (δAux R Q x) := by + induction' x using MvPolynomial.induction_on' with n r x₁ x₂ hx₁ hx₂ + · induction' y using MvPolynomial.induction_on' with m s y₁ y₂ hy₁ hy₂ + · simp only [monomial_mul, δAux_monomial, Derivation.leibniz, tmul_add, tmul_smul, + smul_tmul', smul_eq_mul, Algebra.smul_def, algebraMap_apply, aeval_monomial, mul_assoc] + rw [mul_comm (m.prod _) (n.prod _)] + simp only [pow_zero, implies_true, pow_add, Finsupp.prod_add_index'] + · simp only [map_add, smul_add, hy₁, hy₂, mul_add, add_smul]; abel + · simp only [add_mul, map_add, hx₁, hx₂, add_smul, smul_add]; abel + +lemma δAux_C (r) : + δAux R Q (C r) = 1 ⊗ₜ D R S r := by + rw [← monomial_zero', δAux_monomial, Finsupp.prod_zero_index] + +lemma δAux_toAlgHom {Q : Generators.{u₁} S T} + {Q' : Generators.{u₃} S T} (f : Hom Q Q') (x) : + δAux R Q' (f.toAlgHom x) = δAux R Q x + Finsupp.linearCombination _ (δAux R Q' ∘ f.val) + (Q.cotangentSpaceBasis.repr ((1 : T) ⊗ₜ[Q.Ring] D S Q.Ring x : _)) := by + letI : AddCommGroup (T ⊗[S] Ω[S⁄R]) := inferInstance + have : IsScalarTower Q.Ring Q.Ring T := IsScalarTower.left _ + induction' x using MvPolynomial.induction_on with s x₁ x₂ hx₁ hx₂ p n IH + · simp [MvPolynomial.algebraMap_eq, δAux_C] + · simp only [map_add, hx₁, hx₂, tmul_add] + rw [add_add_add_comm] + · simp only [map_mul, Hom.toAlgHom_X, δAux_mul, algebraMap_apply, Hom.algebraMap_toAlgHom, + ← @IsScalarTower.algebraMap_smul Q'.Ring T, id.map_eq_id, δAux_X, RingHomCompTriple.comp_eq, + RingHom.id_apply, coe_eval₂Hom, IH, Hom.aeval_val, smul_add, map_aeval, tmul_add, tmul_smul, + ← @IsScalarTower.algebraMap_smul Q.Ring T, smul_zero, aeval_X, zero_add, Derivation.leibniz, + LinearEquiv.map_add, LinearEquiv.map_smul, Basis.repr_self, LinearMap.map_add, one_smul, + LinearMap.map_smul, Finsupp.linearCombination_single, + Function.comp_apply, ← cotangentSpaceBasis_apply] + rw [add_left_comm] + rfl + +lemma δAux_ofComp (x : (Q.comp P).Ring) : + δAux R Q ((Q.ofComp P).toAlgHom x) = + P.toExtension.toKaehler.baseChange T (CotangentSpace.compEquiv Q P + (1 ⊗ₜ[(Q.comp P).Ring] (D R (Q.comp P).Ring) x : _)).2 := by + letI : AddCommGroup (T ⊗[S] Ω[S⁄R]) := inferInstance + have : IsScalarTower (Q.comp P).Ring (Q.comp P).Ring T := IsScalarTower.left _ + induction' x using MvPolynomial.induction_on with s x₁ x₂ hx₁ hx₂ p n IH + · simp only [algHom_C, δAux_C, sub_self, derivation_C, Derivation.map_algebraMap, + tmul_zero, map_zero, add_zero, MvPolynomial.algebraMap_apply, Prod.snd_zero] + · simp only [map_add, hx₁, hx₂, tmul_add, Prod.snd_add] + · simp only [map_mul, Hom.toAlgHom_X, ofComp_val, δAux_mul, + ← @IsScalarTower.algebraMap_smul Q.Ring T, algebraMap_apply, Hom.algebraMap_toAlgHom, + id.map_eq_id, map_aeval, RingHomCompTriple.comp_eq, comp_val, RingHom.id_apply, coe_eval₂Hom, + IH, Derivation.leibniz, tmul_add, tmul_smul, ← cotangentSpaceBasis_apply, + ← @IsScalarTower.algebraMap_smul (Q.comp P).Ring T, aeval_X, LinearEquiv.map_add, + LinearMapClass.map_smul, Prod.snd_add, Prod.smul_snd, LinearMap.map_add] + obtain (n | n) := n + · simp only [comp_vars, Sum.elim_inl, δAux_X, smul_zero, aeval_X, + CotangentSpace.compEquiv, LinearEquiv.trans_apply, Basis.repr_symm_apply, zero_add, + Basis.repr_self, Finsupp.linearCombination_single, Basis.prod_apply, LinearMap.coe_inl, + LinearMap.coe_inr, Function.comp_apply, one_smul, map_zero] + · simp only [comp_vars, Sum.elim_inr, Function.comp_apply, algHom_C, δAux_C, + CotangentSpace.compEquiv, LinearEquiv.trans_apply, Basis.repr_symm_apply, + algebraMap_smul, Basis.repr_self, Finsupp.linearCombination_single, Basis.prod_apply, + LinearMap.coe_inr, Basis.baseChange_apply, one_smul, LinearMap.baseChange_tmul, + toKaehler_cotangentSpaceBasis, add_left_inj, LinearMap.coe_inl] + rfl + +lemma map_comp_cotangentComplex_baseChange : + (Extension.CotangentSpace.map (Q.toComp P).toExtensionHom).liftBaseChange T ∘ₗ + P.toExtension.cotangentComplex.baseChange T = + (Q.comp P).toExtension.cotangentComplex ∘ₗ + (Extension.Cotangent.map (Q.toComp P).toExtensionHom).liftBaseChange T := by + ext x; simp [Extension.CotangentSpace.map_cotangentComplex] + +open Generators in +/-- +The connecting homomorphism in the Jacobi-Zariski sequence for given presentations. +Given representations `0 → I → R[X] → S → 0` and `0 → K → S[Y] → T → 0`, +we may consider the induced representation `0 → J → R[X, Y] → T → 0`, +and this map is obtained by applying snake lemma to the following diagram +``` + T ⊗[S] Ω[S/R] → Ω[T/R] → Ω[T/S] → 0 + ↑ ↑ ↑ +0 → T ⊗[S] (⨁ₓ S dx) → (⨁ₓ T dx) ⊕ (⨁ᵧ T dy) → ⨁ᵧ T dy → 0 + ↑ ↑ ↑ + T ⊗[S] (I/I²) → J/J² → K/K² → 0 + ↑ ↑ + H¹(L_{T/R}) → H¹(L_{T/S}) + +``` +This is independent from the presentations chosen. See `H1Cotangent.δ_comp_equiv`. +-/ +noncomputable +def δ : + Q.toExtension.H1Cotangent →ₗ[T] T ⊗[S] Ω[S⁄R] := + SnakeLemma.δ' + (P.toExtension.cotangentComplex.baseChange T) + (Q.comp P).toExtension.cotangentComplex + Q.toExtension.cotangentComplex + ((Extension.Cotangent.map (toComp Q P).toExtensionHom).liftBaseChange T) + (Extension.Cotangent.map (ofComp Q P).toExtensionHom) + (Cotangent.exact Q P) + ((Extension.CotangentSpace.map (toComp Q P).toExtensionHom).liftBaseChange T) + (Extension.CotangentSpace.map (ofComp Q P).toExtensionHom) + (CotangentSpace.exact Q P) + (map_comp_cotangentComplex_baseChange Q P) + (by ext; exact Extension.CotangentSpace.map_cotangentComplex (ofComp Q P).toExtensionHom _) + Q.toExtension.h1Cotangentι + (LinearMap.exact_subtype_ker_map _) + (N₁ := T ⊗[S] P.toExtension.CotangentSpace) + (P.toExtension.toKaehler.baseChange T) + (lTensor_exact T P.toExtension.exact_cotangentComplex_toKaehler + P.toExtension.toKaehler_surjective) + (Cotangent.surjective_map_ofComp Q P) + (CotangentSpace.map_toComp_injective Q P) + +lemma exact_δ_map : + Function.Exact (δ Q P) (mapBaseChange R S T) := by + apply SnakeLemma.exact_δ_left (π₂ := (Q.comp P).toExtension.toKaehler) + (hπ₂ := (Q.comp P).toExtension.exact_cotangentComplex_toKaehler) + · apply (P.cotangentSpaceBasis.baseChange T).ext + intro i + simp only [Basis.baseChange_apply, LinearMap.coe_comp, Function.comp_apply, + LinearMap.baseChange_tmul, toKaehler_cotangentSpaceBasis, mapBaseChange_tmul, map_D, + one_smul, comp_vars, LinearMap.liftBaseChange_tmul] + rw [cotangentSpaceBasis_apply] + conv_rhs => enter [2]; tactic => exact Extension.CotangentSpace.map_tmul .. + simp only [map_one, mapBaseChange_tmul, map_D, one_smul] + simp [Extension.Hom.toAlgHom] + · exact LinearMap.lTensor_surjective T P.toExtension.toKaehler_surjective + +lemma δ_eq (x : Q.toExtension.H1Cotangent) (y) + (hy : Extension.Cotangent.map (ofComp Q P).toExtensionHom y = x.1) (z) + (hz : (Extension.CotangentSpace.map (toComp Q P).toExtensionHom).liftBaseChange T z = + (Q.comp P).toExtension.cotangentComplex y) : + δ Q P x = P.toExtension.toKaehler.baseChange T z := by + apply SnakeLemma.δ_eq + exacts [hy, hz] + +lemma δ_eq_δAux (x : Q.ker) (hx) : + δ Q P ⟨.mk x, hx⟩ = δAux R Q x.1 := by + let y := Extension.Cotangent.mk (P := (Q.comp P).toExtension) (Q.kerCompPreimage P x) + have hy : (Extension.Cotangent.map (Q.ofComp P).toExtensionHom) y = Extension.Cotangent.mk x := by + simp only [y, Extension.Cotangent.map_mk] + congr + exact ofComp_kerCompPreimage Q P x + let z := (CotangentSpace.compEquiv Q P ((Q.comp P).toExtension.cotangentComplex y)).2 + rw [H1Cotangent.δ_eq (y := y) (z := z)] + · rw [← ofComp_kerCompPreimage Q P x, δAux_ofComp] + rfl + · exact hy + · rw [← CotangentSpace.compEquiv_symm_inr] + apply (CotangentSpace.compEquiv Q P).injective + simp only [LinearMap.coe_comp, LinearEquiv.coe_coe, LinearMap.coe_inr, Function.comp_apply, + LinearEquiv.apply_symm_apply, z] + ext + swap; · rfl + show 0 = (LinearMap.fst T Q.toExtension.CotangentSpace (T ⊗[S] P.toExtension.CotangentSpace) ∘ₗ + (CotangentSpace.compEquiv Q P).toLinearMap) ((Q.comp P).toExtension.cotangentComplex y) + rw [CotangentSpace.fst_compEquiv, Extension.CotangentSpace.map_cotangentComplex, hy, hx] + +lemma δ_eq_δ (Q : Generators.{u₁} S T) (P : Generators.{u₂} R S) + (P' : Generators.{u₃} R S) : + δ Q P = δ Q P' := by + ext ⟨x, hx⟩ + obtain ⟨x, rfl⟩ := Extension.Cotangent.mk_surjective x + rw [δ_eq_δAux, δ_eq_δAux] + +lemma exact_map_δ : + Function.Exact (Extension.H1Cotangent.map (Q.ofComp P).toExtensionHom) (δ Q P) := by + apply SnakeLemma.exact_δ_right + (ι₂ := (Q.comp P).toExtension.h1Cotangentι) + (hι₂ := LinearMap.exact_subtype_ker_map _) + · ext x; rfl + · exact Subtype.val_injective + +lemma δ_map + (Q : Generators.{u₁} S T) (P : Generators.{u₂} R S) + (Q' : Generators.{u₃} S T) (P' : Generators.{u₄} R S) (f : Hom Q' Q) (x) : + δ Q P (Extension.H1Cotangent.map f.toExtensionHom x) = δ Q' P' x := by + letI : AddCommGroup (T ⊗[S] Ω[S⁄R]) := inferInstance + obtain ⟨x, hx⟩ := x + obtain ⟨⟨y, hy⟩, rfl⟩ := Extension.Cotangent.mk_surjective x + show δ _ _ ⟨_, _⟩ = δ _ _ _ + replace hx : (1 : T) ⊗ₜ[Q'.Ring] (D S Q'.Ring) y = 0 := by + simpa only [LinearMap.mem_ker, Extension.cotangentComplex_mk, ker, RingHom.mem_ker] using hx + simp only [LinearMap.domRestrict_apply, Extension.Cotangent.map_mk, δ_eq_δAux] + refine (δAux_toAlgHom f _).trans ?_ + rw [hx, map_zero, map_zero, add_zero] + +lemma δ_comp_equiv + (Q : Generators.{u₁} S T) (P : Generators.{u₂} R S) + (Q' : Generators.{u₃} S T) (P' : Generators.{u₄} R S) : + δ Q P ∘ₗ (H1Cotangent.equiv _ _).toLinearMap = δ Q' P' := by + ext x + exact δ_map Q P Q' P' _ _ + +/-- A variant of `exact_map_δ` that takes in an arbitrary map between generators. -/ +lemma exact_map_δ' + (Q : Generators.{u₁} S T) (P : Generators.{u₂} R S) (P' : Generators.{u₃} R T) (f : Hom P' Q) : + Function.Exact (Extension.H1Cotangent.map f.toExtensionHom) (δ Q P) := by + refine (H1Cotangent.equiv (Q.comp P) P').surjective.comp_exact_iff_exact.mp ?_ + show Function.Exact ((Extension.H1Cotangent.map f.toExtensionHom).restrictScalars T ∘ₗ + (Extension.H1Cotangent.map _)) (δ Q P) + rw [← Extension.H1Cotangent.map_comp, Extension.H1Cotangent.map_eq _ (Q.ofComp P).toExtensionHom] + exact exact_map_δ Q P + +end H1Cotangent + +end instanceProblem + +end Generators + +variable {T : Type w} [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] + +variable (R S T) + +/-- The connecting homomorphism in the Jacobi-Zariski sequence. -/ +noncomputable +def H1Cotangent.δ : H1Cotangent S T →ₗ[T] T ⊗[S] Ω[S⁄R] := + Generators.H1Cotangent.δ (Generators.self S T) (Generators.self R S) + +/-- Given algebras `R → S → T`, `H¹(L_{T/R}) → H¹(L_{T/S}) → T ⊗[S] Ω[S/R]` is exact. -/ +lemma H1Cotangent.exact_map_δ : Function.Exact (map R S T T) (δ R S T) := + Generators.H1Cotangent.exact_map_δ' (Generators.self S T) + (Generators.self R S) (Generators.self R T) (Generators.defaultHom _ _) + +/-- Given algebras `R → S → T`, `H¹(L_{T/S}) → T ⊗[S] Ω[S/R] → Ω[T/R]` is exact. -/ +lemma H1Cotangent.exact_δ_mapBaseChange : Function.Exact (δ R S T) (mapBaseChange R S T) := + Generators.H1Cotangent.exact_δ_map (Generators.self S T) (Generators.self R S) + +end Algebra From 32b1ec0957158906fe045e95e16a71d93cd161a0 Mon Sep 17 00:00:00 2001 From: Hannah Scholz Date: Sat, 28 Dec 2024 13:28:22 +0000 Subject: [PATCH 022/189] feat: set is subset of product of its image under projections (#20076) Co-authored-by: Hannah Scholz <70071345+scholzhannah@users.noreply.github.com> --- Mathlib/Data/Set/Prod.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Data/Set/Prod.lean b/Mathlib/Data/Set/Prod.lean index 16e65038f06b0..2aec0e84f338a 100644 --- a/Mathlib/Data/Set/Prod.lean +++ b/Mathlib/Data/Set/Prod.lean @@ -366,6 +366,9 @@ theorem prod_eq_iff_eq (ht : t.Nonempty) : s ×ˢ t = s₁ ×ˢ t ↔ s = s₁ : rintro ⟨rfl, rfl⟩ rfl +theorem subset_prod {s : Set (α × β)} : s ⊆ (Prod.fst '' s) ×ˢ (Prod.snd '' s) := + fun _ hp ↦ mem_prod.2 ⟨mem_image_of_mem _ hp, mem_image_of_mem _ hp⟩ + section Mono variable [Preorder α] {f : α → Set β} {g : α → Set γ} From 0c419e32e8cc2b54b3ac18eafa3595cd17659e1e Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sat, 28 Dec 2024 14:09:27 +0000 Subject: [PATCH 023/189] feat: add `Submonoid.mem_closure_range_iff` (#19495) Co-authored-by: Kim Morrison --- Mathlib.lean | 2 + Mathlib/Algebra/BigOperators/Finsupp.lean | 7 ++- Mathlib/Algebra/Group/Subgroup/Finsupp.lean | 56 ++++++++++++++++++++ Mathlib/Algebra/Group/Submonoid/Finsupp.lean | 49 +++++++++++++++++ Mathlib/Data/Finsupp/Multiset.lean | 2 +- 5 files changed, 114 insertions(+), 2 deletions(-) create mode 100644 Mathlib/Algebra/Group/Subgroup/Finsupp.lean create mode 100644 Mathlib/Algebra/Group/Submonoid/Finsupp.lean diff --git a/Mathlib.lean b/Mathlib.lean index c0bbcc04a8b3b..d6a597218d32d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -317,6 +317,7 @@ import Mathlib.Algebra.Group.Subgroup.Actions import Mathlib.Algebra.Group.Subgroup.Basic import Mathlib.Algebra.Group.Subgroup.Defs import Mathlib.Algebra.Group.Subgroup.Finite +import Mathlib.Algebra.Group.Subgroup.Finsupp import Mathlib.Algebra.Group.Subgroup.Ker import Mathlib.Algebra.Group.Subgroup.Lattice import Mathlib.Algebra.Group.Subgroup.Map @@ -330,6 +331,7 @@ import Mathlib.Algebra.Group.Submonoid.Basic import Mathlib.Algebra.Group.Submonoid.BigOperators import Mathlib.Algebra.Group.Submonoid.Defs import Mathlib.Algebra.Group.Submonoid.DistribMulAction +import Mathlib.Algebra.Group.Submonoid.Finsupp import Mathlib.Algebra.Group.Submonoid.Membership import Mathlib.Algebra.Group.Submonoid.MulOpposite import Mathlib.Algebra.Group.Submonoid.Operations diff --git a/Mathlib/Algebra/BigOperators/Finsupp.lean b/Mathlib/Algebra/BigOperators/Finsupp.lean index fa5e090a17199..056c1427d57a0 100644 --- a/Mathlib/Algebra/BigOperators/Finsupp.lean +++ b/Mathlib/Algebra/BigOperators/Finsupp.lean @@ -115,11 +115,16 @@ theorem sum_ite_self_eq' [DecidableEq α] {N : Type*} [AddCommMonoid N] (f : α convert f.sum_ite_eq' a fun _ => id simp [ite_eq_right_iff.2 Eq.symm] -@[simp] +@[to_additive (attr := simp)] theorem prod_pow [Fintype α] (f : α →₀ ℕ) (g : α → N) : (f.prod fun a b => g a ^ b) = ∏ a, g a ^ f a := f.prod_fintype _ fun _ ↦ pow_zero _ +@[to_additive (attr := simp)] +theorem prod_zpow {N} [CommGroup N] [Fintype α] (f : α →₀ ℤ) (g : α → N) : + (f.prod fun a b => g a ^ b) = ∏ a, g a ^ f a := + f.prod_fintype _ fun _ ↦ zpow_zero _ + /-- If `g` maps a second argument of 0 to 1, then multiplying it over the result of `onFinset` is the same as multiplying it over the original `Finset`. -/ @[to_additive diff --git a/Mathlib/Algebra/Group/Subgroup/Finsupp.lean b/Mathlib/Algebra/Group/Subgroup/Finsupp.lean new file mode 100644 index 0000000000000..50360ca190d27 --- /dev/null +++ b/Mathlib/Algebra/Group/Subgroup/Finsupp.lean @@ -0,0 +1,56 @@ +/- +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.Algebra.BigOperators.Finsupp +import Mathlib.Algebra.Group.Subgroup.Lattice + +/-! # Connection between `Subgroup.closure` and `Finsupp.prod` -/ + +namespace Subgroup + +variable {M : Type*} [CommGroup M] {ι : Type*} (f : ι → M) (x : M) + +@[to_additive] +theorem exists_finsupp_of_mem_closure_range (hx : x ∈ closure (Set.range f)) : + ∃ a : ι →₀ ℤ, x = a.prod (f · ^ ·) := by + classical + induction hx using closure_induction with + | mem x h => obtain ⟨i, rfl⟩ := h; exact ⟨Finsupp.single i 1, by simp⟩ + | one => use 0; simp + | mul x y hx hy hx' hy' => + obtain ⟨⟨v, rfl⟩, w, rfl⟩ := And.intro hx' hy' + use v + w + rw [Finsupp.prod_add_index] + · simp + · simp [zpow_add] + | inv x hx hx' => + obtain ⟨a, rfl⟩ := hx' + use -a + rw [Finsupp.prod_neg_index] + · simp + · simp [zpow_neg] + +variable {f x} in +@[to_additive] +theorem mem_closure_range_iff : + x ∈ closure (Set.range f) ↔ ∃ a : ι →₀ ℤ, x = a.prod (f · ^ ·) := by + refine ⟨exists_finsupp_of_mem_closure_range f x, ?_⟩ + rintro ⟨a, rfl⟩ + exact Submonoid.prod_mem _ fun i hi ↦ zpow_mem (subset_closure (Set.mem_range_self i)) _ + +@[to_additive] +theorem exists_of_mem_closure_range [Fintype ι] (hx : x ∈ closure (Set.range f)) : + ∃ a : ι → ℤ, x = ∏ i, f i ^ a i := by + obtain ⟨a, rfl⟩ := exists_finsupp_of_mem_closure_range f x hx + exact ⟨a, by simp⟩ + +variable {f x} in +@[to_additive] +theorem mem_closure_range_iff_of_fintype [Fintype ι] : + x ∈ closure (Set.range f) ↔ ∃ a : ι → ℤ, x = ∏ i, f i ^ a i := by + rw [Finsupp.equivFunOnFinite.symm.exists_congr_left, mem_closure_range_iff] + simp + +end Subgroup diff --git a/Mathlib/Algebra/Group/Submonoid/Finsupp.lean b/Mathlib/Algebra/Group/Submonoid/Finsupp.lean new file mode 100644 index 0000000000000..838df5cc55165 --- /dev/null +++ b/Mathlib/Algebra/Group/Submonoid/Finsupp.lean @@ -0,0 +1,49 @@ +/- +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.Algebra.BigOperators.Finsupp + +/-! # Connection between `Submonoid.closure` and `Finsupp.prod` -/ + +namespace Submonoid + +variable {M : Type*} [CommMonoid M] {ι : Type*} (f : ι → M) (x : M) + +@[to_additive] +theorem exists_finsupp_of_mem_closure_range (hx : x ∈ closure (Set.range f)) : + ∃ a : ι →₀ ℕ, x = a.prod (f · ^ ·) := by + classical + induction hx using closure_induction with + | mem x h => obtain ⟨i, rfl⟩ := h; exact ⟨Finsupp.single i 1, by simp⟩ + | one => use 0; simp + | mul x y hx hy hx' hy' => + obtain ⟨⟨v, rfl⟩, w, rfl⟩ := And.intro hx' hy' + use v + w + rw [Finsupp.prod_add_index] + · simp + · simp [pow_add] + +variable {f x} in +@[to_additive] +theorem mem_closure_range_iff : + x ∈ closure (Set.range f) ↔ ∃ a : ι →₀ ℕ, x = a.prod (f · ^ ·) := by + refine ⟨exists_finsupp_of_mem_closure_range f x, ?_⟩ + rintro ⟨a, rfl⟩ + exact prod_mem _ fun i hi ↦ pow_mem (subset_closure (Set.mem_range_self i)) _ + +@[to_additive] +theorem exists_of_mem_closure_range [Fintype ι] (hx : x ∈ closure (Set.range f)) : + ∃ a : ι → ℕ, x = ∏ i, f i ^ a i := by + obtain ⟨a, rfl⟩ := exists_finsupp_of_mem_closure_range f x hx + exact ⟨a, by simp⟩ + +variable {f x} in +@[to_additive] +theorem mem_closure_range_iff_of_fintype [Fintype ι] : + x ∈ closure (Set.range f) ↔ ∃ a : ι → ℕ, x = ∏ i, f i ^ a i := by + rw [Finsupp.equivFunOnFinite.symm.exists_congr_left, mem_closure_range_iff] + simp + +end Submonoid diff --git a/Mathlib/Data/Finsupp/Multiset.lean b/Mathlib/Data/Finsupp/Multiset.lean index 39a01e2fb9894..df6435246cdea 100644 --- a/Mathlib/Data/Finsupp/Multiset.lean +++ b/Mathlib/Data/Finsupp/Multiset.lean @@ -53,7 +53,7 @@ theorem toMultiset_sum {f : ι → α →₀ ℕ} (s : Finset ι) : theorem toMultiset_sum_single (s : Finset ι) (n : ℕ) : Finsupp.toMultiset (∑ i ∈ s, single i n) = n • s.val := by - simp_rw [toMultiset_sum, Finsupp.toMultiset_single, sum_nsmul, sum_multiset_singleton] + simp_rw [toMultiset_sum, Finsupp.toMultiset_single, Finset.sum_nsmul, sum_multiset_singleton] @[simp] theorem card_toMultiset (f : α →₀ ℕ) : Multiset.card (toMultiset f) = f.sum fun _ => id := by From 1bf90b6e1851c4ff7722cfb149bd0ac0295995f7 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Sat, 28 Dec 2024 18:52:52 +0000 Subject: [PATCH 024/189] feat(GroupTheory/SpecificGroups/ZGroup): A finite Z-group has cyclic commutator subgroup (#20183) One more step towards classifying finite Z-groups. --- Mathlib/Algebra/Group/Hom/Basic.lean | 10 ++++++ Mathlib/Algebra/Group/Subgroup/Map.lean | 5 +++ Mathlib/GroupTheory/Abelianization.lean | 5 +++ .../GroupTheory/SpecificGroups/ZGroup.lean | 36 +++++++++++++++++++ 4 files changed, 56 insertions(+) diff --git a/Mathlib/Algebra/Group/Hom/Basic.lean b/Mathlib/Algebra/Group/Hom/Basic.lean index 114dff59a75ba..aa6909aaf8420 100644 --- a/Mathlib/Algebra/Group/Hom/Basic.lean +++ b/Mathlib/Algebra/Group/Hom/Basic.lean @@ -228,4 +228,14 @@ lemma comp_div (f : G →* H) (g h : M →* G) : f.comp (g / h) = f.comp g / f.c end InvDiv +/-- If `H` is commutative and `G →* H` is injective, then `G` is commutative. -/ +def commGroupOfInjective [Group G] [CommGroup H] (f : G →* H) (hf : Function.Injective f) : + CommGroup G := + ⟨by simp_rw [← hf.eq_iff, map_mul, mul_comm, implies_true]⟩ + +/-- If `G` is commutative and `G →* H` is surjective, then `H` is commutative. -/ +def commGroupOfSurjective [CommGroup G] [Group H] (f : G →* H) (hf : Function.Surjective f) : + CommGroup H := + ⟨by simp_rw [hf.forall₂, ← map_mul, mul_comm, implies_true]⟩ + end MonoidHom diff --git a/Mathlib/Algebra/Group/Subgroup/Map.lean b/Mathlib/Algebra/Group/Subgroup/Map.lean index e1219dcd59203..76fdb39866cc6 100644 --- a/Mathlib/Algebra/Group/Subgroup/Map.lean +++ b/Mathlib/Algebra/Group/Subgroup/Map.lean @@ -298,6 +298,11 @@ theorem mem_subgroupOf {H K : Subgroup G} {h : K} : h ∈ H.subgroupOf K ↔ (h theorem subgroupOf_map_subtype (H K : Subgroup G) : (H.subgroupOf K).map K.subtype = H ⊓ K := SetLike.ext' <| by refine Subtype.image_preimage_coe _ _ |>.trans ?_; apply Set.inter_comm +@[to_additive] +theorem map_subgroupOf_eq_of_le {H K : Subgroup G} (h : H ≤ K) : + (H.subgroupOf K).map K.subtype = H := by + rwa [subgroupOf_map_subtype, inf_eq_left] + @[to_additive (attr := simp)] theorem bot_subgroupOf : (⊥ : Subgroup G).subgroupOf H = ⊥ := Eq.symm (Subgroup.ext fun _g => Subtype.ext_iff) diff --git a/Mathlib/GroupTheory/Abelianization.lean b/Mathlib/GroupTheory/Abelianization.lean index e27a628bddbad..cc6ce0148ce0d 100644 --- a/Mathlib/GroupTheory/Abelianization.lean +++ b/Mathlib/GroupTheory/Abelianization.lean @@ -50,6 +50,11 @@ theorem commutator_eq_closure : commutator G = Subgroup.closure (commutatorSet G theorem commutator_eq_normalClosure : commutator G = Subgroup.normalClosure (commutatorSet G) := by simp [commutator, Subgroup.commutator_def', commutatorSet] +variable {G} in +theorem Subgroup.map_subtype_commutator (H : Subgroup G) : + (_root_.commutator H).map H.subtype = ⁅H, H⁆ := by + rw [_root_.commutator_def, map_commutator, ← MonoidHom.range_eq_map, H.range_subtype] + instance commutator_characteristic : (commutator G).Characteristic := Subgroup.commutator_characteristic ⊤ ⊤ diff --git a/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean b/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean index f9e5473f66bf2..84f690ccc8fc3 100644 --- a/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean +++ b/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean @@ -19,6 +19,7 @@ A Z-group is a group whose Sylow subgroups are all cyclic. ## Main results * `IsZGroup.isCyclic_abelianization`: a finite Z-group has cyclic abelianization. +* `IsZGroup.isCyclic_commutator`: a finite Z-group has cyclic commutator subgroup. TODO: Show that if `G` is a Z-group with commutator subgroup `G'`, then `G = G' ⋊ G/G'` where `G'` and `G/G'` are cyclic of coprime orders. @@ -126,6 +127,41 @@ instance isCyclic_abelianization [Finite G] [IsZGroup G] : IsCyclic (Abelianizat end Nilpotent +section Commutator + +variable (G) in +/-- A finite Z-group has cyclic commutator subgroup. -/ +theorem isCyclic_commutator [Finite G] [IsZGroup G] : IsCyclic (commutator G) := by + refine WellFoundedLT.induction (C := fun H ↦ IsCyclic (⁅H, H⁆ : Subgroup G)) (⊤ : Subgroup G) ?_ + intro H hH + rcases eq_or_ne H ⊥ with rfl | h + · rw [Subgroup.commutator_bot_left] + infer_instance + · specialize hH ⁅H, H⁆ (IsSolvable.commutator_lt_of_ne_bot h) + replace hH : IsCyclic (⁅commutator H, commutator H⁆ : Subgroup H) := by + let f := Subgroup.equivMapOfInjective ⁅commutator H, commutator H⁆ _ H.subtype_injective + rw [Subgroup.map_commutator, Subgroup.map_subtype_commutator] at f + exact isCyclic_of_surjective f.symm f.symm.surjective + suffices IsCyclic (commutator H) by + let f := Subgroup.equivMapOfInjective (commutator H) _ H.subtype_injective + rw [Subgroup.map_subtype_commutator] at f + exact isCyclic_of_surjective f f.surjective + suffices h : commutator (commutator H) ≤ Subgroup.center (commutator H) by + rw [← Abelianization.ker_of (commutator H)] at h + let _ := commGroupOfCyclicCenterQuotient Abelianization.of h + infer_instance + suffices h : (commutator (commutator H)).map (commutator H).subtype ≤ + Subgroup.centralizer (commutator H) by + simpa [SetLike.le_def, Subgroup.mem_center_iff, Subgroup.mem_centralizer_iff] using h + rw [Subgroup.map_subtype_commutator, Subgroup.le_centralizer_iff] + let _ := (hH.mulAutMulEquiv _).toMonoidHom.commGroupOfInjective (hH.mulAutMulEquiv _).injective + have h := Abelianization.commutator_subset_ker ⁅commutator H, commutator H⁆.normalizerMonoidHom + rwa [Subgroup.normalizerMonoidHom_ker, Subgroup.normalizer_eq_top, + ← Subgroup.map_subtype_le_map_subtype, Subgroup.map_subtype_commutator, + Subgroup.map_subgroupOf_eq_of_le le_top] at h + +end Commutator + section Classification /-- An extension of coprime Z-groups is a Z-group. -/ From 1a402974fd696ae8df2990028b9fac0752700d79 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 28 Dec 2024 23:51:51 +0000 Subject: [PATCH 025/189] chore: rename `update_same`/`update_noteq` to `update_self`/`update_of_ne` (#20302) The previous names didn't match the naming convention --- Mathlib/Algebra/Group/Pi/Basic.lean | 4 +- Mathlib/Algebra/Group/Pi/Lemmas.lean | 2 +- .../Group/Pointwise/Set/BigOperators.lean | 6 +- Mathlib/Algebra/Group/Subgroup/Finite.lean | 2 +- Mathlib/Algebra/Group/TypeTags/Basic.lean | 4 +- Mathlib/Algebra/Order/Antidiag/Finsupp.lean | 6 +- Mathlib/Analysis/Analytic/Composition.lean | 4 +- .../BoxIntegral/Partition/Additive.lean | 6 +- .../Calculus/ContDiff/FaaDiBruno.lean | 24 +++---- Mathlib/Analysis/Calculus/DSlope.lean | 4 +- .../Analysis/Calculus/FDeriv/Analytic.lean | 8 +-- Mathlib/Analysis/Complex/CauchyIntegral.lean | 2 +- Mathlib/Analysis/Complex/Periodic.lean | 10 +-- Mathlib/Analysis/Convex/Extreme.lean | 8 +-- .../NormedSpace/Multilinear/Basic.lean | 2 +- .../PiTensorProduct/InjectiveSeminorm.lean | 4 +- Mathlib/CategoryTheory/Filtered/Final.lean | 4 +- Mathlib/Computability/TMToPartrec.lean | 36 +++++----- Mathlib/Computability/TuringMachine.lean | 18 ++--- Mathlib/Data/DFinsupp/Defs.lean | 6 +- Mathlib/Data/Fin/Tuple/Basic.lean | 6 +- Mathlib/Data/Fin/Tuple/Take.lean | 6 +- Mathlib/Data/Finset/PiInduction.lean | 6 +- Mathlib/Data/Finset/Piecewise.lean | 4 +- Mathlib/Data/Finset/Update.lean | 4 +- Mathlib/Data/Finsupp/Defs.lean | 4 +- Mathlib/Data/Fintype/Prod.lean | 2 +- Mathlib/Data/Matrix/RowCol.lean | 12 ++-- Mathlib/Data/Nat/Choose/Multinomial.lean | 10 +-- Mathlib/Data/Set/Card.lean | 2 +- Mathlib/Data/Set/Lattice.lean | 2 +- Mathlib/Data/Set/Prod.lean | 8 +-- Mathlib/Data/Sigma/Basic.lean | 6 +- Mathlib/Data/Sum/Basic.lean | 4 +- Mathlib/Geometry/Manifold/Instances/Real.lean | 8 +-- Mathlib/GroupTheory/Complement.lean | 12 ++-- .../AffineSpace/Combination.lean | 6 +- .../AffineSpace/Independent.lean | 4 +- Mathlib/LinearAlgebra/Alternating/Basic.lean | 4 +- Mathlib/LinearAlgebra/Matrix/Basis.lean | 4 +- Mathlib/LinearAlgebra/Multilinear/Basic.lean | 66 +++++++++---------- Mathlib/LinearAlgebra/Pi.lean | 4 +- Mathlib/LinearAlgebra/PiTensorProduct.lean | 4 +- Mathlib/Logic/Equiv/Basic.lean | 5 +- Mathlib/Logic/Function/Basic.lean | 20 +++--- Mathlib/MeasureTheory/Constructions/Pi.lean | 4 +- .../Integral/FundThmCalculus.lean | 4 +- .../Integral/IntegralEqImproper.lean | 8 +-- Mathlib/MeasureTheory/Integral/Marginal.lean | 2 +- .../LSeries/DirichletContinuation.lean | 8 +-- .../NumberTheory/LSeries/HurwitzZetaEven.lean | 26 ++++---- .../NumberTheory/LSeries/Nonvanishing.lean | 10 +-- Mathlib/NumberTheory/LSeries/PrimesInAP.lean | 2 +- Mathlib/NumberTheory/LSeries/RiemannZeta.lean | 4 +- Mathlib/NumberTheory/LSeries/ZMod.lean | 2 +- .../CanonicalEmbedding/ConvexBody.lean | 6 +- Mathlib/Order/Atoms.lean | 10 +-- Mathlib/Order/Basic.lean | 4 +- Mathlib/Order/Bounds/Image.lean | 2 +- Mathlib/Order/CompleteLattice/Finset.lean | 4 +- Mathlib/Order/Interval/Set/Pi.lean | 8 +-- Mathlib/Order/Lattice.lean | 4 +- Mathlib/Order/PiLex.lean | 18 ++--- Mathlib/RingTheory/Ideal/Maps.lean | 4 +- Mathlib/RingTheory/Localization/Basic.lean | 6 +- Mathlib/RingTheory/MaximalSpectrum.lean | 8 +-- Mathlib/RingTheory/MvPowerSeries/Basic.lean | 6 +- Mathlib/RingTheory/PrimeSpectrum.lean | 2 +- .../Topology/Algebra/InfiniteSum/Group.lean | 4 +- .../Algebra/Module/Multilinear/Bounded.lean | 4 +- .../Topology/Algebra/WithZeroTopology.lean | 2 +- Mathlib/Topology/ContinuousOn.lean | 4 +- Mathlib/Topology/Order.lean | 4 +- Mathlib/Topology/Separation/Basic.lean | 12 ++-- Mathlib/Topology/ShrinkingLemma.lean | 30 ++++----- 75 files changed, 294 insertions(+), 291 deletions(-) diff --git a/Mathlib/Algebra/Group/Pi/Basic.lean b/Mathlib/Algebra/Group/Pi/Basic.lean index 14f4dc1890c74..9016e7a5abc7f 100644 --- a/Mathlib/Algebra/Group/Pi/Basic.lean +++ b/Mathlib/Algebra/Group/Pi/Basic.lean @@ -259,11 +259,11 @@ def mulSingle (i : I) (x : f i) : ∀ (j : I), f j := @[to_additive (attr := simp)] theorem mulSingle_eq_same (i : I) (x : f i) : mulSingle i x i = x := - Function.update_same i x _ + Function.update_self i x _ @[to_additive (attr := simp)] theorem mulSingle_eq_of_ne {i i' : I} (h : i' ≠ i) (x : f i) : mulSingle i x i' = 1 := - Function.update_noteq h x _ + Function.update_of_ne h x _ /-- Abbreviation for `mulSingle_eq_of_ne h.symm`, for ease of use by `simp`. -/ @[to_additive (attr := simp) diff --git a/Mathlib/Algebra/Group/Pi/Lemmas.lean b/Mathlib/Algebra/Group/Pi/Lemmas.lean index 0f412767df4f9..1d2924faf728b 100644 --- a/Mathlib/Algebra/Group/Pi/Lemmas.lean +++ b/Mathlib/Algebra/Group/Pi/Lemmas.lean @@ -289,7 +289,7 @@ theorem Pi.update_eq_div_mul_mulSingle [∀ i, Group <| f i] (g : ∀ i : I, f i ext j rcases eq_or_ne i j with (rfl | h) · simp - · simp [Function.update_noteq h.symm, h] + · simp [Function.update_of_ne h.symm, h] @[to_additive] theorem Pi.mulSingle_mul_mulSingle_eq_mulSingle_mul_mulSingle {M : Type*} [CommMonoid M] diff --git a/Mathlib/Algebra/Group/Pointwise/Set/BigOperators.lean b/Mathlib/Algebra/Group/Pointwise/Set/BigOperators.lean index 7772e8e35548f..c4675e6c8534d 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/BigOperators.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/BigOperators.lean @@ -62,10 +62,10 @@ theorem mem_finset_prod (t : Finset ι) (f : ι → Set α) (a : α) : refine ⟨Function.update g i x, ?_, ?_⟩ · intro j hj obtain rfl | hj := Finset.mem_insert.mp hj - · rwa [Function.update_same] - · rw [update_noteq (ne_of_mem_of_not_mem hj hi)] + · rwa [Function.update_self] + · rw [update_of_ne (ne_of_mem_of_not_mem hj hi)] exact hg hj - · rw [Finset.prod_update_of_not_mem hi, Function.update_same] + · rw [Finset.prod_update_of_not_mem hi, Function.update_self] · rintro ⟨g, hg, rfl⟩ exact ⟨g i, hg (is.mem_insert_self _), is.prod g, ⟨⟨g, fun hi ↦ hg (Finset.mem_insert_of_mem hi), rfl⟩, rfl⟩⟩ diff --git a/Mathlib/Algebra/Group/Subgroup/Finite.lean b/Mathlib/Algebra/Group/Subgroup/Finite.lean index 05d21ec805804..092fb9835528b 100644 --- a/Mathlib/Algebra/Group/Subgroup/Finite.lean +++ b/Mathlib/Algebra/Group/Subgroup/Finite.lean @@ -192,7 +192,7 @@ theorem pi_mem_of_mulSingle_mem_aux [DecidableEq η] (I : Finset η) {H : Subgro have : j ≠ i := by rintro rfl contradiction - simp only [ne_eq, this, not_false_eq_true, Function.update_noteq] + simp only [ne_eq, this, not_false_eq_true, Function.update_of_ne] exact h2 _ (Finset.mem_insert_of_mem hj) · apply h2 simp diff --git a/Mathlib/Algebra/Group/TypeTags/Basic.lean b/Mathlib/Algebra/Group/TypeTags/Basic.lean index 152313ad2ebc6..a856c3ad37498 100644 --- a/Mathlib/Algebra/Group/TypeTags/Basic.lean +++ b/Mathlib/Algebra/Group/TypeTags/Basic.lean @@ -448,7 +448,7 @@ lemma Pi.mulSingle_multiplicativeOfAdd_eq {ι : Type*} [DecidableEq ι] {M : ι Multiplicative.ofAdd ((Pi.single i a) j) := by rcases eq_or_ne j i with rfl | h · simp only [mulSingle_eq_same, single_eq_same] - · simp only [mulSingle, ne_eq, h, not_false_eq_true, Function.update_noteq, one_apply, single, + · simp only [mulSingle, ne_eq, h, not_false_eq_true, Function.update_of_ne, one_apply, single, zero_apply, ofAdd_zero] lemma Pi.single_additiveOfMul_eq {ι : Type*} [DecidableEq ι] {M : ι → Type*} @@ -457,5 +457,5 @@ lemma Pi.single_additiveOfMul_eq {ι : Type*} [DecidableEq ι] {M : ι → Type* Additive.ofMul ((Pi.mulSingle i a) j) := by rcases eq_or_ne j i with rfl | h · simp only [mulSingle_eq_same, single_eq_same] - · simp only [single, ne_eq, h, not_false_eq_true, Function.update_noteq, zero_apply, mulSingle, + · simp only [single, ne_eq, h, not_false_eq_true, Function.update_of_ne, zero_apply, mulSingle, one_apply, ofMul_one] diff --git a/Mathlib/Algebra/Order/Antidiag/Finsupp.lean b/Mathlib/Algebra/Order/Antidiag/Finsupp.lean index 150ff45ab21d7..0f45c2dc2ee32 100644 --- a/Mathlib/Algebra/Order/Antidiag/Finsupp.lean +++ b/Mathlib/Algebra/Order/Antidiag/Finsupp.lean @@ -68,7 +68,7 @@ theorem mem_finsuppAntidiag_insert {a : ι} {s : Finset ι} constructor · rintro ⟨rfl, hsupp⟩ refine ⟨_, _, rfl, Finsupp.erase a f, ?_, ?_, ?_⟩ - · rw [update_erase_eq_update, update_self] + · rw [update_erase_eq_update, Finsupp.update_self] · apply sum_congr rfl intro x hx rw [Finsupp.erase_ne (ne_of_mem_of_not_mem hx h)] @@ -77,10 +77,10 @@ theorem mem_finsuppAntidiag_insert {a : ι} {s : Finset ι} refine ⟨?_, (support_update_subset _ _).trans (insert_subset_insert a hgsupp)⟩ simp only [coe_update] apply congr_arg₂ - · rw [update_same] + · rw [Function.update_self] · apply sum_congr rfl intro x hx - rw [update_noteq (ne_of_mem_of_not_mem hx h) n1 ⇑g] + rw [update_of_ne (ne_of_mem_of_not_mem hx h) n1 ⇑g] theorem finsuppAntidiag_insert {a : ι} {s : Finset ι} (h : a ∉ s) (n : μ) : diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 29001e428114c..8badf38fc79a0 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -141,14 +141,14 @@ theorem applyComposition_update (p : FormalMultilinearSeries 𝕜 E F) {n : ℕ} by_cases h : k = c.index j · rw [h] let r : Fin (c.blocksFun (c.index j)) → Fin n := c.embedding (c.index j) - simp only [Function.update_same] + simp only [Function.update_self] change p (c.blocksFun (c.index j)) (Function.update v j z ∘ r) = _ let j' := c.invEmbedding j suffices B : Function.update v j z ∘ r = Function.update (v ∘ r) j' z by rw [B] suffices C : Function.update v (r j') z ∘ r = Function.update (v ∘ r) j' z by convert C; exact (c.embedding_comp_inv j).symm exact Function.update_comp_eq_of_injective _ (c.embedding _).injective _ _ - · simp only [h, Function.update_eq_self, Function.update_noteq, Ne, not_false_iff] + · simp only [h, Function.update_eq_self, Function.update_of_ne, Ne, not_false_iff] let r : Fin (c.blocksFun k) → Fin n := c.embedding k change p (c.blocksFun k) (Function.update v j z ∘ r) = p (c.blocksFun k) (v ∘ r) suffices B : Function.update v j z ∘ r = v ∘ r by rw [B] diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean b/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean index f82d7bce340fe..adb77e996ca75 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean @@ -188,8 +188,8 @@ def upperSubLower.{u} {G : Type u} [AddCommGroup G] (I₀ : Box (Fin (n + 1))) ( intro J hJ j x rw [WithTop.coe_le_coe] at hJ refine i.succAboveCases (fun hx => ?_) (fun j hx => ?_) j - · simp only [Box.splitLower_def hx, Box.splitUpper_def hx, update_same, ← WithBot.some_eq_coe, - Option.elim', Box.face, Function.comp_def, update_noteq (Fin.succAbove_ne _ _)] + · simp only [Box.splitLower_def hx, Box.splitUpper_def hx, update_self, ← WithBot.some_eq_coe, + Option.elim', Box.face, Function.comp_def, update_of_ne (Fin.succAbove_ne _ _)] abel · have : (J.face i : WithTop (Box (Fin n))) ≤ I₀.face i := WithTop.coe_le_coe.2 (face_mono hJ i) @@ -200,7 +200,7 @@ def upperSubLower.{u} {G : Type u} [AddCommGroup G] (I₀ : Box (Fin (n + 1))) ( have hx' : x ∈ Ioo ((J.face i).lower j) ((J.face i).upper j) := hx simp only [Box.splitLower_def hx, Box.splitUpper_def hx, Box.splitLower_def hx', Box.splitUpper_def hx', ← WithBot.some_eq_coe, Option.elim', Box.face_mk, - update_noteq (Fin.succAbove_ne _ _).symm, sub_add_sub_comm, + update_of_ne (Fin.succAbove_ne _ _).symm, sub_add_sub_comm, update_comp_eq_of_injective _ (Fin.strictMono_succAbove i).injective j x, ← hf] simp only [Box.face]) diff --git a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean index c42683bf68966..347525ce0df81 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean @@ -357,7 +357,7 @@ def extendMiddle (c : OrderedFinpartition n) (k : Fin c.length) : OrderedFinpart parts_strictMono := by convert strictMono_succ.comp c.parts_strictMono with m rcases eq_or_ne m k with rfl | hm - · simp only [↓reduceDIte, update_same, add_tsub_cancel_right, comp_apply, cast_mk, + · simp only [↓reduceDIte, update_self, add_tsub_cancel_right, comp_apply, cast_mk, Nat.succ_eq_add_one] let a : Fin (c.partSize m + 1) := ⟨c.partSize m, lt_add_one (c.partSize m)⟩ let b : Fin (c.partSize m) := ⟨c.partSize m - 1, Nat.sub_one_lt_of_lt (c.partSize_pos m)⟩ @@ -493,7 +493,7 @@ def eraseMiddle (c : OrderedFinpartition (n + 1)) (hc : range (c.emb 0) ≠ {0}) partSize_pos i := by rcases eq_or_ne i (c.index 0) with rfl | hi · simpa using c.one_lt_partSize_index_zero hc - · simp only [ne_eq, hi, not_false_eq_true, update_noteq] + · simp only [ne_eq, hi, not_false_eq_true, update_of_ne] exact c.partSize_pos i emb i j := by by_cases h : i = c.index 0 @@ -517,7 +517,7 @@ def eraseMiddle (c : OrderedFinpartition (n + 1)) (hc : range (c.emb 0) ≠ {0}) rw [← Nat.add_lt_add_iff_right (k := 1)] convert Fin.lt_iff_val_lt_val.1 (c.parts_strictMono hij) · rcases eq_or_ne i (c.index 0) with rfl | hi - · simp only [↓reduceDIte, Nat.succ_eq_add_one, update_same, succ_mk, cast_mk, coe_pred] + · simp only [↓reduceDIte, Nat.succ_eq_add_one, update_self, succ_mk, cast_mk, coe_pred] have A := c.one_lt_partSize_index_zero hc rw [Nat.sub_add_cancel] · congr; omega @@ -526,14 +526,14 @@ def eraseMiddle (c : OrderedFinpartition (n + 1)) (hc : range (c.emb 0) ≠ {0}) rw [← lt_iff_val_lt_val] apply c.emb_strictMono simp [lt_iff_val_lt_val] - · simp only [hi, ↓reduceDIte, ne_eq, not_false_eq_true, update_noteq, cast_mk, coe_pred] + · simp only [hi, ↓reduceDIte, ne_eq, not_false_eq_true, update_of_ne, cast_mk, coe_pred] apply Nat.sub_add_cancel have : c.emb i ⟨c.partSize i - 1, Nat.sub_one_lt_of_lt (c.partSize_pos i)⟩ ≠ c.emb (c.index 0) 0 := c.emb_ne_emb_of_ne hi simp only [c.emb_zero, ne_eq, ← val_eq_val, val_zero] at this omega · rcases eq_or_ne j (c.index 0) with rfl | hj - · simp only [↓reduceDIte, Nat.succ_eq_add_one, update_same, succ_mk, cast_mk, coe_pred] + · simp only [↓reduceDIte, Nat.succ_eq_add_one, update_self, succ_mk, cast_mk, coe_pred] have A := c.one_lt_partSize_index_zero hc rw [Nat.sub_add_cancel] · congr; omega @@ -542,7 +542,7 @@ def eraseMiddle (c : OrderedFinpartition (n + 1)) (hc : range (c.emb 0) ≠ {0}) rw [← lt_iff_val_lt_val] apply c.emb_strictMono simp [lt_iff_val_lt_val] - · simp only [hj, ↓reduceDIte, ne_eq, not_false_eq_true, update_noteq, cast_mk, coe_pred] + · simp only [hj, ↓reduceDIte, ne_eq, not_false_eq_true, update_of_ne, cast_mk, coe_pred] apply Nat.sub_add_cancel have : c.emb j ⟨c.partSize j - 1, Nat.sub_one_lt_of_lt (c.partSize_pos j)⟩ ≠ c.emb (c.index 0) 0 := c.emb_ne_emb_of_ne hj @@ -667,12 +667,12 @@ def extendEquiv (n : ℕ) : simp only [extend, extendMiddle, eraseMiddle, Nat.succ_eq_add_one, ↓reduceDIte] ext · rfl - · simp only [update_same, update_idem, heq_eq_eq, update_eq_self_iff, B] + · simp only [update_self, update_idem, heq_eq_eq, update_eq_self_iff, B] · refine hfunext rfl ?_ simp only [heq_eq_eq, forall_eq'] intro i refine ((Fin.heq_fun_iff ?_).mpr ?_).symm - · simp only [update_same, B, update_idem, update_eq_self] + · simp only [update_self, B, update_idem, update_eq_self] · intro j rcases eq_or_ne i (c.index 0) with rfl | hi · simp only [↓reduceDIte, comp_apply] @@ -717,13 +717,13 @@ theorem applyOrderedFinpartition_update_right ext m by_cases h : m = c.index j · rw [h] - simp only [applyOrderedFinpartition, update_same] + simp only [applyOrderedFinpartition, update_self] congr rw [← Function.update_comp_eq_of_injective] · simp · exact (c.emb_strictMono (c.index j)).injective · simp only [applyOrderedFinpartition, ne_eq, h, not_false_eq_true, - update_noteq] + update_of_ne] congr apply Function.update_comp_eq_of_not_mem_range have A : Disjoint (range (c.emb m)) (range (c.emb (c.index j))) := @@ -913,14 +913,14 @@ private lemma faaDiBruno_aux2 {m : ℕ} (q : FormalMultilinearSeries 𝕜 F G) congr ext j rcases eq_or_ne j i with rfl | hij - · simp only [↓reduceDIte, update_same, ContinuousMultilinearMap.curryLeft_apply, + · simp only [↓reduceDIte, update_self, ContinuousMultilinearMap.curryLeft_apply, Nat.succ_eq_add_one] apply FormalMultilinearSeries.congr _ (by simp) intro a ha h'a match a with | 0 => simp | a + 1 => simp [cons] - · simp only [hij, ↓reduceDIte, ne_eq, not_false_eq_true, update_noteq] + · simp only [hij, ↓reduceDIte, ne_eq, not_false_eq_true, update_of_ne] apply FormalMultilinearSeries.congr _ (by simp [hij]) simp diff --git a/Mathlib/Analysis/Calculus/DSlope.lean b/Mathlib/Analysis/Calculus/DSlope.lean index df304274fc2ac..144d4924f8fea 100644 --- a/Mathlib/Analysis/Calculus/DSlope.lean +++ b/Mathlib/Analysis/Calculus/DSlope.lean @@ -34,13 +34,13 @@ noncomputable def dslope (f : 𝕜 → E) (a : 𝕜) : 𝕜 → E := @[simp] theorem dslope_same (f : 𝕜 → E) (a : 𝕜) : dslope f a a = deriv f a := by classical - exact update_same _ _ _ + exact update_self .. variable {f : 𝕜 → E} {a b : 𝕜} {s : Set 𝕜} theorem dslope_of_ne (f : 𝕜 → E) (h : b ≠ a) : dslope f a b = slope f a b := by classical - exact update_noteq h _ _ + exact update_of_ne 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) : diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index a0f9f214a88b5..59fd0722065d8 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -595,8 +595,8 @@ theorem changeOrigin_toFormalMultilinearSeries [DecidableEq ι] : Function.update_apply, (Equiv.injective _).eq_iff, ite_apply] congr; ext j obtain rfl | hj := eq_or_ne j i - · rw [Function.update_same, if_pos rfl] - · rw [Function.update_noteq hj, if_neg hj] + · rw [Function.update_self, if_pos rfl] + · rw [Function.update_of_ne hj, if_neg hj] protected theorem hasFDerivAt [DecidableEq ι] : HasFDerivAt f (f.linearDeriv x) x := by rw [← changeOrigin_toFormalMultilinearSeries] @@ -688,7 +688,7 @@ theorem hasFTaylorSeriesUpTo_iteratedFDeriv : exact Set.range_comp_subset_range _ _ hke simp only [hke, hkf, ↓reduceDIte, Pi.compRightL, ContinuousLinearMap.coe_mk', LinearMap.coe_mk, AddHom.coe_mk] - rw [Function.update_noteq] + rw [Function.update_of_ne] contrapose! hke rw [show k = _ from Subtype.ext_iff_val.1 hke, Equiv.embeddingFinSucc_snd e] exact Set.mem_range_self _ @@ -812,7 +812,7 @@ theorem hasFDerivAt_uncurry_of_multilinear [DecidableEq ι] simp only [ContinuousMultilinearMap.toContinuousLinearMap, continuousMultilinearMapOption, coe_mk', MultilinearMap.toLinearMap_apply, ContinuousMultilinearMap.coe_coe, MultilinearMap.coe_mkContinuous, MultilinearMap.coe_mk, ne_eq, reduceCtorEq, - not_false_eq_true, Function.update_noteq, coe_comp', coe_snd', Function.comp_apply, + not_false_eq_true, Function.update_of_ne, coe_comp', coe_snd', Function.comp_apply, proj_apply] congr ext j diff --git a/Mathlib/Analysis/Complex/CauchyIntegral.lean b/Mathlib/Analysis/Complex/CauchyIntegral.lean index 19be8544d83fe..99b3e0d3c22dc 100644 --- a/Mathlib/Analysis/Complex/CauchyIntegral.lean +++ b/Mathlib/Analysis/Complex/CauchyIntegral.lean @@ -427,7 +427,7 @@ theorem circleIntegral_sub_inv_smul_of_differentiable_on_off_countable_aux {R : have hne : ∀ z ∈ sphere c R, z ≠ w := fun z hz => ne_of_mem_of_not_mem hz (ne_of_lt hw.1) have hFeq : EqOn F (fun z => (z - w)⁻¹ • f z - (z - w)⁻¹ • f w) (sphere c R) := fun z hz ↦ calc - F z = (z - w)⁻¹ • (f z - f w) := update_noteq (hne z hz) _ _ + F z = (z - w)⁻¹ • (f z - f w) := update_of_ne (hne z hz) .. _ = (z - w)⁻¹ • f z - (z - w)⁻¹ • f w := smul_sub _ _ _ have hc' : ContinuousOn (fun z => (z - w)⁻¹) (sphere c R) := (continuousOn_id.sub continuousOn_const).inv₀ fun z hz => sub_ne_zero.2 <| hne z hz diff --git a/Mathlib/Analysis/Complex/Periodic.lean b/Mathlib/Analysis/Complex/Periodic.lean index d8c7535894d28..d0c7d17b29266 100644 --- a/Mathlib/Analysis/Complex/Periodic.lean +++ b/Mathlib/Analysis/Complex/Periodic.lean @@ -96,20 +96,20 @@ def cuspFunction : ℂ → ℂ := theorem cuspFunction_eq_of_nonzero {q : ℂ} (hq : q ≠ 0) : cuspFunction h f q = f (invQParam h q) := - update_noteq hq .. + update_of_ne hq .. theorem cuspFunction_zero_eq_limUnder_nhds_ne : cuspFunction h f 0 = limUnder (𝓝[≠] 0) (cuspFunction h f) := by - conv_lhs => simp only [cuspFunction, update_same] + conv_lhs => simp only [cuspFunction, update_self] refine congr_arg lim (Filter.map_congr <| eventuallyEq_nhdsWithin_of_eqOn fun r hr ↦ ?_) - rw [cuspFunction, update_noteq hr] + rw [cuspFunction, update_of_ne hr] variable {f h} theorem eq_cuspFunction (hh : h ≠ 0) (hf : Periodic f h) (z : ℂ) : (cuspFunction h f) (𝕢 h z) = f z := by have : (cuspFunction h f) (𝕢 h z) = f (invQParam h (𝕢 h z)) := by - rw [cuspFunction, update_noteq, comp_apply] + rw [cuspFunction, update_of_ne, comp_apply] exact exp_ne_zero _ obtain ⟨m, hm⟩ := qParam_left_inv_mod_period hh z simpa only [this, hm] using hf.int_mul m z @@ -169,7 +169,7 @@ theorem boundedAtFilter_cuspFunction (hh : 0 < h) (h_bd : BoundedAtFilter I∞ f theorem cuspFunction_zero_of_zero_at_inf (hh : 0 < h) (h_zer : ZeroAtFilter I∞ f) : cuspFunction h f 0 = 0 := by - simpa only [cuspFunction, update_same] using (h_zer.comp (invQParam_tendsto hh)).limUnder_eq + simpa only [cuspFunction, update_self] using (h_zer.comp (invQParam_tendsto hh)).limUnder_eq theorem differentiableAt_cuspFunction_zero (hh : 0 < h) (hf : Periodic f h) (h_hol : ∀ᶠ z in I∞, DifferentiableAt ℂ f z) (h_bd : BoundedAtFilter I∞ f) : diff --git a/Mathlib/Analysis/Convex/Extreme.lean b/Mathlib/Analysis/Convex/Extreme.lean index a59c44b680d98..652e14f30bcc9 100644 --- a/Mathlib/Analysis/Convex/Extreme.lean +++ b/Mathlib/Analysis/Convex/Extreme.lean @@ -191,13 +191,13 @@ theorem extremePoints_pi (s : ∀ i, Set (π i)) : simp only [mem_extremePoints, mem_pi, mem_univ, true_imp_iff, @forall_and ι] refine and_congr_right fun hx ↦ ⟨fun h i ↦ ?_, fun h ↦ ?_⟩ · rintro x₁ hx₁ x₂ hx₂ hi - refine (h (update x i x₁) ?_ (update x i x₂) ?_ ?_).imp (fun h₁ ↦ by rw [← h₁, update_same]) - fun h₂ ↦ by rw [← h₂, update_same] + refine (h (update x i x₁) ?_ (update x i x₂) ?_ ?_).imp (fun h₁ ↦ by rw [← h₁, update_self]) + fun h₂ ↦ by rw [← h₂, update_self] iterate 2 rintro j obtain rfl | hji := eq_or_ne j i - · rwa [update_same] - · rw [update_noteq hji] + · rwa [update_self] + · rw [update_of_ne hji] exact hx _ rw [← Pi.image_update_openSegment] exact ⟨_, hi, update_eq_self _ _⟩ diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index ca6f026c41a7c..96a773a054737 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -262,7 +262,7 @@ theorem norm_image_sub_le_of_bound (f : MultilinearMap 𝕜 E G) · intro j _ by_cases h : j = i · rw [h] - simp only [ite_true, Function.update_same] + simp only [ite_true, Function.update_self] exact norm_le_pi_norm (m₁ - m₂) i · simp [h, - le_sup_iff, - sup_le_iff, sup_le_sup, norm_le_pi_norm] _ = ‖m₁ - m₂‖ * max ‖m₁‖ ‖m₂‖ ^ (Fintype.card ι - 1) := by diff --git a/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean b/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean index 7b0fa337cc06a..b8ce28e14f8a2 100644 --- a/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean +++ b/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean @@ -411,8 +411,8 @@ private theorem mapL_add_smul_aux {ι : Type uι} symm rw [update_eq_iff] constructor - · simp only [update_same] - · exact fun _ h ↦ by simp only [ne_eq, h, not_false_eq_true, update_noteq] + · simp only [update_self] + · exact fun _ h ↦ by simp only [ne_eq, h, not_false_eq_true, update_of_ne] open Function in protected theorem mapL_add [DecidableEq ι] (i : ι) (u v : E i →L[𝕜] E' i) : diff --git a/Mathlib/CategoryTheory/Filtered/Final.lean b/Mathlib/CategoryTheory/Filtered/Final.lean index 8710c29ec9105..98c6d31fcbbdb 100644 --- a/Mathlib/CategoryTheory/Filtered/Final.lean +++ b/Mathlib/CategoryTheory/Filtered/Final.lean @@ -395,7 +395,7 @@ instance final_eval [∀ s, IsFiltered (I s)] (s : α) : (Pi.eval I s).Final := s ⟨coeq f g, coeqHom f g⟩ refine ⟨fun t => (c't t).1, fun t => (c't t).2, ?_⟩ dsimp only [Pi.eval_obj, Pi.eval_map, c't] - rw [Function.update_same] + rw [Function.update_self] simpa using coeq_condition _ _ open IsCofiltered in @@ -408,7 +408,7 @@ instance initial_eval [∀ s, IsCofiltered (I s)] (s : α) : (Pi.eval I s).Initi s ⟨eq f g, eqHom f g⟩ refine ⟨fun t => (c't t).1, fun t => (c't t).2, ?_⟩ dsimp only [Pi.eval_obj, Pi.eval_map, c't] - rw [Function.update_same] + rw [Function.update_self] simpa using eq_condition _ _ end Pi diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index 0b08e16e8fbb1..232291b05955f 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -1236,7 +1236,7 @@ theorem move_ok {p k₁ k₂ q s L₁ o L₂} {S : K' → List Γ'} (h₁ : k₁ induction' L₁ with a L₁ IH generalizing S s · rw [(_ : [].reverseAux _ = _), Function.update_eq_self] swap - · rw [Function.update_noteq h₁.symm, List.reverseAux_nil] + · rw [Function.update_of_ne h₁.symm, List.reverseAux_nil] refine TransGen.head' rfl ?_ simp only [TM2.step, Option.mem_def, TM2.stepAux, Option.elim, ne_eq] revert e; cases' S k₁ with a Sk <;> intro e @@ -1259,7 +1259,7 @@ theorem move_ok {p k₁ k₂ q s L₁ o L₂} {S : K' → List Γ'} (h₁ : k₁ cases e simp only [List.head?_cons, e₂, List.tail_cons, ne_eq, cond_false] convert @IH _ (update (update S k₁ Sk) k₂ (a :: S k₂)) _ using 2 <;> - simp [Function.update_noteq, h₁, h₁.symm, e₃, List.reverseAux] + simp [Function.update_of_ne, h₁, h₁.symm, e₃, List.reverseAux] simp [Function.update_comm h₁.symm] theorem unrev_ok {q s} {S : K' → List Γ'} : @@ -1278,16 +1278,16 @@ theorem move₂_ok {p k₁ k₂ q s L₁ o L₂} {S : K' → List Γ'} (h₁ : k convert move_ok h₁.2.1.symm (splitAtPred_false _) using 2 simp only [Function.update_comm h₁.1, Function.update_idem] rw [show update S rev [] = S by rw [← h₂, Function.update_eq_self]] - simp only [Function.update_noteq h₁.2.2.symm, Function.update_noteq h₁.2.1, - Function.update_noteq h₁.1.symm, List.reverseAux_eq, h₂, Function.update_same, + simp only [Function.update_of_ne h₁.2.2.symm, Function.update_of_ne h₁.2.1, + Function.update_of_ne h₁.1.symm, List.reverseAux_eq, h₂, Function.update_self, List.append_nil, List.reverse_reverse] · simp only [TM2.stepAux, Option.isSome, cond_true] convert move_ok h₁.2.1.symm (splitAtPred_false _) using 2 - simp only [h₂, Function.update_comm h₁.1, List.reverseAux_eq, Function.update_same, + simp only [h₂, Function.update_comm h₁.1, List.reverseAux_eq, Function.update_self, List.append_nil, Function.update_idem] rw [show update S rev [] = S by rw [← h₂, Function.update_eq_self]] - simp only [Function.update_noteq h₁.1.symm, Function.update_noteq h₁.2.2.symm, - Function.update_noteq h₁.2.1, Function.update_same, List.reverse_reverse] + simp only [Function.update_of_ne h₁.1.symm, Function.update_of_ne h₁.2.2.symm, + Function.update_of_ne h₁.2.1, Function.update_self, List.reverse_reverse] theorem clear_ok {p k q s L₁ o L₂} {S : K' → List Γ'} (e : splitAtPred p (S k) = (L₁, o, L₂)) : Reaches₁ (TM2.step tr) ⟨some (Λ'.clear p k q), s, S⟩ ⟨some q, o, update S k L₂⟩ := by @@ -1323,7 +1323,7 @@ theorem copy_ok (q s a b c d) : simp refine TransGen.head rfl ?_ simp only [TM2.step, Option.mem_def, TM2.stepAux, elim_rev, List.head?_cons, Option.isSome_some, - List.tail_cons, elim_update_rev, ne_eq, Function.update_noteq, elim_main, elim_update_main, + List.tail_cons, elim_update_rev, ne_eq, Function.update_of_ne, elim_main, elim_update_main, elim_stack, elim_update_stack, cond_true, List.reverseAux_cons] exact IH _ _ _ @@ -1359,7 +1359,7 @@ theorem head_main_ok {q s L} {c d : List Γ'} : (TransGen.head rfl (TransGen.head rfl ?_)) · cases L <;> simp [o] simp only [TM2.step, Option.mem_def, TM2.stepAux, elim_update_main, elim_rev, elim_update_rev, - Function.update_same, trList] + Function.update_self, trList] rw [if_neg (show o ≠ some Γ'.consₗ by cases L <;> simp [o])] refine (clear_ok (splitAtPred_eq _ _ _ none [] ?_ ⟨rfl, rfl⟩)).trans ?_ · exact fun x h => Bool.decide_false (trList_ne_consₗ _ _ h) @@ -1376,7 +1376,7 @@ theorem head_stack_ok {q s L₁ L₂ L₃} : (splitAtPred_eq _ _ [] (some Γ'.consₗ) L₃ (by rintro _ ⟨⟩) ⟨rfl, rfl⟩)) (TransGen.head rfl (TransGen.head rfl ?_)) simp only [TM2.step, Option.mem_def, TM2.stepAux, ite_true, id_eq, trList, List.nil_append, - elim_update_stack, elim_rev, List.reverseAux_nil, elim_update_rev, Function.update_same, + elim_update_stack, elim_rev, List.reverseAux_nil, elim_update_rev, Function.update_self, List.headI_nil, trNat_default] convert unrev_ok using 2 simp @@ -1387,7 +1387,7 @@ theorem head_stack_ok {q s L₁ L₂ L₃} : (trNat_natEnd _) ⟨rfl, by simp⟩)) (TransGen.head rfl (TransGen.head rfl ?_)) simp only [TM2.step, Option.mem_def, TM2.stepAux, ite_false, trList, List.append_assoc, - List.cons_append, elim_update_stack, elim_rev, elim_update_rev, Function.update_same, + List.cons_append, elim_update_stack, elim_rev, elim_update_rev, Function.update_self, List.headI_cons] refine TransGen.trans @@ -1405,7 +1405,7 @@ theorem succ_ok {q s n} {c d : List Γ'} : cases' (n : Num) with a · refine TransGen.head rfl ?_ simp only [Option.mem_def, TM2.stepAux, elim_main, decide_false, elim_update_main, ne_eq, - Function.update_noteq, elim_rev, elim_update_rev, decide_true, Function.update_same, + Function.update_of_ne, elim_rev, elim_update_rev, decide_true, Function.update_self, cond_true, cond_false] convert unrev_ok using 1 simp only [elim_update_rev, elim_rev, elim_main, List.reverseAux_nil, elim_update_main] @@ -1428,8 +1428,8 @@ theorem succ_ok {q s n} {c d : List Γ'} : simp [PosNum.succ, trPosNum] rfl · refine ⟨l₁, _, some Γ'.bit0, rfl, TransGen.single ?_⟩ - simp only [TM2.step, TM2.stepAux, elim_main, elim_update_main, ne_eq, Function.update_noteq, - elim_rev, elim_update_rev, Function.update_same, Option.mem_def, Option.some.injEq] + simp only [TM2.step, TM2.stepAux, elim_main, elim_update_main, ne_eq, Function.update_of_ne, + elim_rev, elim_update_rev, Function.update_self, Option.mem_def, Option.some.injEq] rfl theorem pred_ok (q₁ q₂ s v) (c d : List Γ') : ∃ s', @@ -1448,8 +1448,8 @@ theorem pred_ok (q₁ q₂ s v) (c d : List Γ') : ∃ s', · simp only [trPosNum, List.singleton_append, List.nil_append] refine TransGen.head rfl ?_ simp only [Option.mem_def, TM2.stepAux, elim_main, List.head?_cons, Option.some.injEq, - decide_false, List.tail_cons, elim_update_main, ne_eq, Function.update_noteq, elim_rev, - elim_update_rev, natEnd, Function.update_same, cond_true, cond_false] + decide_false, List.tail_cons, elim_update_main, ne_eq, Function.update_of_ne, elim_rev, + elim_update_rev, natEnd, Function.update_self, cond_true, cond_false] convert unrev_ok using 2 simp simp only [Num.succ'] @@ -1496,7 +1496,7 @@ theorem trNormal_respects (c k v s) : obtain ⟨c, h₁, h₂⟩ := IHf (Cont.cons₁ fs v k) v none refine ⟨c, h₁, TransGen.head rfl <| (move_ok (by decide) (splitAtPred_false _)).trans ?_⟩ simp only [TM2.step, Option.mem_def, elim_stack, elim_update_stack, elim_update_main, ne_eq, - Function.update_noteq, elim_main, elim_rev, elim_update_rev] + Function.update_of_ne, elim_main, elim_rev, elim_update_rev] refine (copy_ok _ none [] (trList v).reverse _ _).trans ?_ convert h₂ using 2 simp [List.reverseAux_eq, trContStack] @@ -1530,7 +1530,7 @@ theorem tr_ret_respects (k v s) : ∃ b₂, (fun x h => Bool.decide_false (trList_ne_consₗ _ _ h)) ⟨rfl, rfl⟩ refine (move₂_ok (by decide) ?_ (splitAtPred_false _)).trans ?_; · rfl simp only [TM2.step, Option.mem_def, Option.elim, elim_update_stack, elim_main, - List.append_nil, elim_update_main, id_eq, elim_update_aux, ne_eq, Function.update_noteq, + List.append_nil, elim_update_main, id_eq, elim_update_aux, ne_eq, Function.update_of_ne, elim_aux, elim_stack] exact h₂ | cons₂ ns k IH => diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index 64621943971af..919e0be38ca9f 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -2305,7 +2305,7 @@ theorem tr_respects_aux₂ [DecidableEq K] {k : K} {q : Stmt₂₁} {v : σ} {S TM1.stepAux (trStAct q o) v ((Tape.move Dir.right)^[(S k).length] (Tape.mk' ∅ (addBottom L))) = TM1.stepAux q v' ((Tape.move Dir.right)^[(S' k).length] (Tape.mk' ∅ (addBottom L'))) := by - simp only [Function.update_same]; cases o with simp only [stWrite, stVar, trStAct, TM1.stepAux] + simp only [Function.update_self]; cases o with simp only [stWrite, stVar, trStAct, TM1.stepAux] | push f => have := Tape.write_move_right_n fun a : Γ' ↦ (a.1, update a.2 k (some (f v))) refine @@ -2320,7 +2320,7 @@ theorem tr_respects_aux₂ [DecidableEq K] {k : K} {q : Stmt₂₁} {v : σ} {S by_cases h' : k' = k · subst k' split_ifs with h - <;> simp only [List.reverse_cons, Function.update_same, ListBlank.nth_mk, List.map] + <;> simp only [List.reverse_cons, Function.update_self, ListBlank.nth_mk, List.map] · rw [List.getI_eq_getElem _, List.getElem_append_right] <;> simp only [List.length_append, List.length_reverse, List.length_map, ← h, Nat.sub_self, List.length_singleton, List.getElem_singleton, @@ -2333,8 +2333,8 @@ theorem tr_respects_aux₂ [DecidableEq K] {k : K} {q : Stmt₂₁} {v : σ} {S rw [List.getI_eq_default, List.getI_eq_default] <;> simp only [Nat.add_one_le_iff, h, List.length, le_of_lt, List.length_reverse, List.length_append, List.length_map] - · split_ifs <;> rw [Function.update_noteq h', ← proj_map_nth, hL] - rw [Function.update_noteq h'] + · split_ifs <;> rw [Function.update_of_ne h', ← proj_map_nth, hL] + rw [Function.update_of_ne h'] | peek f => rw [Function.update_eq_self] use L, hL; rw [Tape.move_left_right]; congr @@ -2363,7 +2363,7 @@ theorem tr_respects_aux₂ [DecidableEq K] {k : K} {q : Stmt₂₁} {v : σ} {S rw [ListBlank.nth_map, ListBlank.nth_modifyNth, proj, PointedMap.mk_val] by_cases h' : k' = k · subst k' - split_ifs with h <;> simp only [Function.update_same, ListBlank.nth_mk, List.tail] + split_ifs with h <;> simp only [Function.update_self, ListBlank.nth_mk, List.tail] · rw [List.getI_eq_default] · rfl rw [h, List.length_reverse, List.length_map] @@ -2375,8 +2375,8 @@ theorem tr_respects_aux₂ [DecidableEq K] {k : K} {q : Stmt₂₁} {v : σ} {S rw [List.getI_eq_default, List.getI_eq_default] <;> simp only [Nat.add_one_le_iff, h, List.length, le_of_lt, List.length_reverse, List.length_append, List.length_map] - · split_ifs <;> rw [Function.update_noteq h', ← proj_map_nth, hL] - rw [Function.update_noteq h'] + · split_ifs <;> rw [Function.update_of_ne h', ← proj_map_nth, hL] + rw [Function.update_of_ne h'] end @@ -2480,9 +2480,9 @@ theorem trCfg_init (k) (L : List (Γ k)) : TrCfg (TM2.init k L) (TM1.init (trIni simp only [] by_cases h : k' = k · subst k' - simp only [Function.update_same] + simp only [Function.update_self] rw [ListBlank.nth_mk, List.getI_eq_iget_getElem?, ← List.map_reverse, List.getElem?_map] - · simp only [Function.update_noteq h] + · simp only [Function.update_of_ne h] rw [ListBlank.nth_mk, List.getI_eq_iget_getElem?, List.map, List.reverse_nil] cases L.reverse[i]? <;> rfl · rw [trInit, TM1.init] diff --git a/Mathlib/Data/DFinsupp/Defs.lean b/Mathlib/Data/DFinsupp/Defs.lean index c6780365cd08d..cdb8f19b4f3a3 100644 --- a/Mathlib/Data/DFinsupp/Defs.lean +++ b/Mathlib/Data/DFinsupp/Defs.lean @@ -622,7 +622,7 @@ def update : Π₀ i, β i := · simp · obtain hj | (hj : f j = 0) := s.prop j · exact Or.inl (Multiset.mem_cons_of_mem hj) - · exact Or.inr ((Function.update_noteq hi.symm b _).trans hj)⟩⟩ + · exact Or.inr ((Function.update_of_ne hi.symm b _).trans hj)⟩⟩ variable (j : ι) @@ -645,14 +645,14 @@ theorem update_eq_single_add_erase {β : ι → Type*} [∀ i, AddZeroClass (β ext j rcases eq_or_ne i j with (rfl | h) · simp - · simp [Function.update_noteq h.symm, h, erase_ne, h.symm] + · simp [Function.update_of_ne h.symm, h, erase_ne, h.symm] theorem update_eq_erase_add_single {β : ι → Type*} [∀ i, AddZeroClass (β i)] (f : Π₀ i, β i) (i : ι) (b : β i) : f.update i b = f.erase i + single i b := by ext j rcases eq_or_ne i j with (rfl | h) · simp - · simp [Function.update_noteq h.symm, h, erase_ne, h.symm] + · simp [Function.update_of_ne h.symm, h, erase_ne, h.symm] theorem update_eq_sub_add_single {β : ι → Type*} [∀ i, AddGroup (β i)] (f : Π₀ i, β i) (i : ι) (b : β i) : f.update i b = f - single i (f i) + single i b := by diff --git a/Mathlib/Data/Fin/Tuple/Basic.lean b/Mathlib/Data/Fin/Tuple/Basic.lean index 430153e9bdb75..72ddaaa0b028e 100644 --- a/Mathlib/Data/Fin/Tuple/Basic.lean +++ b/Mathlib/Data/Fin/Tuple/Basic.lean @@ -135,7 +135,7 @@ theorem cons_update : cons x (update p i y) = update (cons x p) i.succ y := by · rw [h'] simp · have : j'.succ ≠ i.succ := by rwa [Ne, succ_inj] - rw [update_noteq h', update_noteq this, cons_succ] + rw [update_of_ne h', update_of_ne this, cons_succ] /-- As a binary function, `Fin.cons` is injective. -/ theorem cons_injective2 : Function.Injective2 (@cons n α) := fun x₀ y₀ x y h ↦ @@ -159,7 +159,7 @@ theorem update_cons_zero : update (cons x p) 0 z = cons z p := by by_cases h : j = 0 · rw [h] simp - · simp only [h, update_noteq, Ne, not_false_iff] + · simp only [h, update_of_ne, Ne, not_false_iff] let j' := pred j h have : j'.succ = j := succ_pred j h rw [← this, cons_succ, cons_succ] @@ -582,7 +582,7 @@ theorem update_snoc_last : update (snoc p x) (last n) z = snoc p z := by ext j by_cases h : j.val < n · have : j ≠ last n := Fin.ne_of_lt h - simp [h, update_noteq, this, snoc] + simp [h, update_of_ne, this, snoc] · rw [eq_last_of_not_lt h] simp diff --git a/Mathlib/Data/Fin/Tuple/Take.lean b/Mathlib/Data/Fin/Tuple/Take.lean index 462351c63a8d1..232622468b153 100644 --- a/Mathlib/Data/Fin/Tuple/Take.lean +++ b/Mathlib/Data/Fin/Tuple/Take.lean @@ -92,9 +92,9 @@ theorem take_update_of_lt (m : ℕ) (h : m ≤ n) (v : (i : Fin n) → α i) (i ext j by_cases h' : j = i · rw [h'] - simp only [take, update_same] + simp only [take, update_self] · have : castLE h j ≠ castLE h i := by simp [h'] - simp only [take, update_noteq h', update_noteq this] + simp only [take, update_of_ne h', update_of_ne this] /-- `take` is the same after `update` for indices outside the range of `take`. -/ @[simp] @@ -105,7 +105,7 @@ theorem take_update_of_ge (m : ℕ) (h : m ≤ n) (v : (i : Fin n) → α i) (i refine ne_of_val_ne ?_ simp only [coe_castLE] exact Nat.ne_of_lt (lt_of_lt_of_le j.isLt hi) - simp only [take, update_noteq this] + simp only [take, update_of_ne this] /-- Taking the first `m ≤ n` elements of an `addCases u v`, where `u` is a `n`-tuple, is the same as taking the first `m` elements of `u`. -/ diff --git a/Mathlib/Data/Finset/PiInduction.lean b/Mathlib/Data/Finset/PiInduction.lean index 36eed999b59bf..dc87ec1f45595 100644 --- a/Mathlib/Data/Finset/PiInduction.lean +++ b/Mathlib/Data/Finset/PiInduction.lean @@ -50,12 +50,12 @@ theorem induction_on_pi_of_choice (r : ∀ i, α i → Finset (α i) → Prop) set g := update f i ((f i).erase x) with hg clear_value g have hx' : x ∉ g i := by - rw [hg, update_same] + rw [hg, update_self] apply not_mem_erase rw [show f = update g i (insert x (g i)) by - rw [hg, update_idem, update_same, insert_erase x_mem, update_eq_self]] at hr ihs ⊢ + rw [hg, update_idem, update_self, insert_erase x_mem, update_eq_self]] at hr ihs ⊢ clear hg - rw [update_same, erase_insert hx'] at hr + rw [update_self, erase_insert hx'] at hr refine step _ _ _ hr (ihs (univ.sigma g) ?_ _ rfl) rw [ssubset_iff_of_subset (sigma_mono (Subset.refl _) _)] exacts [⟨⟨i, x⟩, mem_sigma.2 ⟨mem_univ _, by simp⟩, by simp [hx']⟩, diff --git a/Mathlib/Data/Finset/Piecewise.lean b/Mathlib/Data/Finset/Piecewise.lean index 99cea5cc557d0..de70ce5c33ede 100644 --- a/Mathlib/Data/Finset/Piecewise.lean +++ b/Mathlib/Data/Finset/Piecewise.lean @@ -105,13 +105,13 @@ lemma update_piecewise [DecidableEq ι] (i : ι) (v : π i) : lemma update_piecewise_of_mem [DecidableEq ι] {i : ι} (hi : i ∈ s) (v : π i) : update (s.piecewise f g) i v = s.piecewise (update f i v) g := by rw [update_piecewise] - refine s.piecewise_congr (fun _ _ => rfl) fun j hj => update_noteq ?_ _ _ + refine s.piecewise_congr (fun _ _ => rfl) fun j hj => update_of_ne ?_ .. exact fun h => hj (h.symm ▸ hi) lemma update_piecewise_of_not_mem [DecidableEq ι] {i : ι} (hi : i ∉ s) (v : π i) : update (s.piecewise f g) i v = s.piecewise f (update g i v) := by rw [update_piecewise] - refine s.piecewise_congr (fun j hj => update_noteq ?_ _ _) fun _ _ => rfl + refine s.piecewise_congr (fun j hj => update_of_ne ?_ ..) fun _ _ => rfl exact fun h => hi (h ▸ hj) lemma piecewise_same : s.piecewise f f = f := by diff --git a/Mathlib/Data/Finset/Update.lean b/Mathlib/Data/Finset/Update.lean index d4dfb055c18ea..53be70e269239 100644 --- a/Mathlib/Data/Finset/Update.lean +++ b/Mathlib/Data/Finset/Update.lean @@ -37,7 +37,7 @@ theorem updateFinset_singleton {i y} : congr with j by_cases hj : j = i · cases hj - simp only [dif_pos, Finset.mem_singleton, update_same, updateFinset] + simp only [dif_pos, Finset.mem_singleton, update_self, updateFinset] · simp [hj, updateFinset] theorem update_eq_updateFinset {i y} : @@ -45,7 +45,7 @@ theorem update_eq_updateFinset {i y} : congr with j by_cases hj : j = i · cases hj - simp only [dif_pos, Finset.mem_singleton, update_same, updateFinset] + simp only [dif_pos, Finset.mem_singleton, update_self, updateFinset] exact uniqueElim_default (α := fun j : ({i} : Finset ι) => π j) y · simp [hj, updateFinset] diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index be9894c31daf0..525b7a364dc57 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -1012,7 +1012,7 @@ theorem update_eq_single_add_erase (f : α →₀ M) (a : α) (b : M) : ext j rcases eq_or_ne a j with (rfl | h) · simp - · simp [Function.update_noteq h.symm, single_apply, h, erase_ne, h.symm] + · simp [Function.update_of_ne h.symm, single_apply, h, erase_ne, h.symm] theorem update_eq_erase_add_single (f : α →₀ M) (a : α) (b : M) : f.update a b = f.erase a + single a b := by @@ -1020,7 +1020,7 @@ theorem update_eq_erase_add_single (f : α →₀ M) (a : α) (b : M) : ext j rcases eq_or_ne a j with (rfl | h) · simp - · simp [Function.update_noteq h.symm, single_apply, h, erase_ne, h.symm] + · simp [Function.update_of_ne h.symm, single_apply, h, erase_ne, h.symm] theorem single_add_erase (a : α) (f : α →₀ M) : single a (f a) + f.erase a = f := by rw [← update_eq_single_add_erase, update_self] diff --git a/Mathlib/Data/Fintype/Prod.lean b/Mathlib/Data/Fintype/Prod.lean index 2a5dfcdbbbf59..882ddcb288b7e 100644 --- a/Mathlib/Data/Fintype/Prod.lean +++ b/Mathlib/Data/Fintype/Prod.lean @@ -71,7 +71,7 @@ instance Pi.infinite_of_left {ι : Sort*} {π : ι → Type*} [∀ i, Nontrivial Infinite (∀ i : ι, π i) := by choose m n hm using fun i => exists_pair_ne (π i) refine Infinite.of_injective (fun i => update m i (n i)) fun x y h => of_not_not fun hne => ?_ - simp_rw [update_eq_iff, update_noteq hne] at h + simp_rw [update_eq_iff, update_of_ne hne] at h exact (hm x h.1.symm).elim /-- If at least one `π i` is infinite and the rest nonempty, the pi type of all `π` is infinite. -/ diff --git a/Mathlib/Data/Matrix/RowCol.lean b/Mathlib/Data/Matrix/RowCol.lean index eabacd0158165..60ca108c68a31 100644 --- a/Mathlib/Data/Matrix/RowCol.lean +++ b/Mathlib/Data/Matrix/RowCol.lean @@ -181,25 +181,25 @@ variable {M : Matrix m n α} {i : m} {j : n} {b : n → α} {c : m → α} @[simp] theorem updateRow_self [DecidableEq m] : updateRow M i b i = b := -- Porting note: (implicit arg) added `(β := _)` - Function.update_same (β := fun _ => (n → α)) i b M + Function.update_self (β := fun _ => (n → α)) i b M @[simp] theorem updateCol_self [DecidableEq n] : updateCol M j c i j = c i := -- Porting note: (implicit arg) added `(β := _)` - Function.update_same (β := fun _ => α) j (c i) (M i) + Function.update_self (β := fun _ => α) j (c i) (M i) @[deprecated (since := "2024-12-11")] alias updateColumn_self := updateCol_self @[simp] theorem updateRow_ne [DecidableEq m] {i' : m} (i_ne : i' ≠ i) : updateRow M i b i' = M i' := -- Porting note: (implicit arg) added `(β := _)` - Function.update_noteq (β := fun _ => (n → α)) i_ne b M + Function.update_of_ne (β := fun _ => (n → α)) i_ne b M @[simp] theorem updateCol_ne [DecidableEq n] {j' : n} (j_ne : j' ≠ j) : updateCol M j c i j' = M i j' := -- Porting note: (implicit arg) added `(β := _)` - Function.update_noteq (β := fun _ => α) j_ne (c i) (M i) + Function.update_of_ne (β := fun _ => α) j_ne (c i) (M i) @[deprecated (since := "2024-12-11")] alias updateColumn_ne := updateCol_ne @@ -288,8 +288,8 @@ theorem diagonal_updateCol_single [DecidableEq n] [Zero α] (v : n → α) (i : obtain rfl | hjk := eq_or_ne j k · rw [diagonal_apply_eq] obtain rfl | hji := eq_or_ne j i - · rw [updateCol_self, Pi.single_eq_same, Function.update_same] - · rw [updateCol_ne hji, diagonal_apply_eq, Function.update_noteq hji] + · rw [updateCol_self, Pi.single_eq_same, Function.update_self] + · rw [updateCol_ne hji, diagonal_apply_eq, Function.update_of_ne hji] · rw [diagonal_apply_ne _ hjk] obtain rfl | hki := eq_or_ne k i · rw [updateCol_self, Pi.single_eq_of_ne hjk] diff --git a/Mathlib/Data/Nat/Choose/Multinomial.lean b/Mathlib/Data/Nat/Choose/Multinomial.lean index 51ba07f30c0c7..318cee9f748cd 100644 --- a/Mathlib/Data/Nat/Choose/Multinomial.lean +++ b/Mathlib/Data/Nat/Choose/Multinomial.lean @@ -121,8 +121,8 @@ theorem binomial_succ_succ [DecidableEq α] (h : a ≠ b) : theorem succ_mul_binomial [DecidableEq α] (h : a ≠ b) : (f a + f b).succ * multinomial {a, b} f = (f a).succ * multinomial {a, b} (Function.update f a (f a).succ) := by - rw [binomial_eq_choose h, binomial_eq_choose h, mul_comm (f a).succ, Function.update_same, - Function.update_noteq (ne_comm.mp h)] + rw [binomial_eq_choose h, binomial_eq_choose h, mul_comm (f a).succ, Function.update_self, + Function.update_of_ne h.symm] rw [succ_mul_choose_eq (f a + f b) (f a), succ_add (f a) (f b)] /-! ### Simple cases -/ @@ -164,7 +164,7 @@ theorem multinomial_update (a : α) (f : α →₀ ℕ) : · rw [← Finset.insert_erase h, Nat.multinomial_insert (Finset.not_mem_erase a _), Finset.add_sum_erase _ f h, support_update_zero] congr 1 - exact Nat.multinomial_congr fun _ h ↦ (Function.update_noteq (mem_erase.1 h).1 0 f).symm + exact Nat.multinomial_congr fun _ h ↦ (Function.update_of_ne (mem_erase.1 h).1 0 f).symm rw [not_mem_support_iff] at h rw [h, Nat.choose_zero_right, one_mul, ← h, update_self] @@ -188,8 +188,8 @@ theorem multinomial_filter_ne [DecidableEq α] (a : α) (m : Multiset α) : · ext1 a rw [toFinsupp_apply, count_filter, Finsupp.coe_update] split_ifs with h - · rw [Function.update_noteq h.symm, toFinsupp_apply] - · rw [not_ne_iff.1 h, Function.update_same] + · rw [Function.update_of_ne h.symm, toFinsupp_apply] + · rw [not_ne_iff.1 h, Function.update_self] @[simp] theorem multinomial_zero [DecidableEq α] : multinomial (0 : Multiset α) = 1 := by diff --git a/Mathlib/Data/Set/Card.lean b/Mathlib/Data/Set/Card.lean index d5b5ef3a80dc0..042625a24b33b 100644 --- a/Mathlib/Data/Set/Card.lean +++ b/Mathlib/Data/Set/Card.lean @@ -436,7 +436,7 @@ theorem Finite.exists_injOn_of_encard_le [Nonempty β] {s : Set α} {t : Set β} · rintro x hx; split_ifs with h · assumption · exact (hf₀s x hx h).1 - exact InjOn.congr hinj (fun x ⟨_, hxa⟩ ↦ by rwa [Function.update_noteq]) + exact InjOn.congr hinj (fun x ⟨_, hxa⟩ ↦ by rwa [Function.update_of_ne]) termination_by encard s theorem Finite.exists_bijOn_of_encard_eq [Nonempty β] (hs : s.Finite) (h : s.encard = t.encard) : diff --git a/Mathlib/Data/Set/Lattice.lean b/Mathlib/Data/Set/Lattice.lean index a47cf4a85b109..0a3b5101f9ede 100644 --- a/Mathlib/Data/Set/Lattice.lean +++ b/Mathlib/Data/Set/Lattice.lean @@ -512,7 +512,7 @@ theorem image_projection_prod {ι : Type*} {α : ι → Type*} {v : ∀ i : ι, · intro y y_in simp only [mem_image, mem_iInter, mem_preimage] rcases hv with ⟨z, hz⟩ - refine ⟨Function.update z i y, ?_, update_same i y z⟩ + refine ⟨Function.update z i y, ?_, update_self i y z⟩ rw [@forall_update_iff ι α _ z i y fun i t => t ∈ v i] exact ⟨y_in, fun j _ => by simpa using hz j⟩ diff --git a/Mathlib/Data/Set/Prod.lean b/Mathlib/Data/Set/Prod.lean index 2aec0e84f338a..742054186e40a 100644 --- a/Mathlib/Data/Set/Prod.lean +++ b/Mathlib/Data/Set/Prod.lean @@ -750,7 +750,7 @@ theorem pi_inter_compl (s : Set ι) : pi s t ∩ pi sᶜ t = pi univ t := by theorem pi_update_of_not_mem [DecidableEq ι] (hi : i ∉ s) (f : ∀ j, α j) (a : α i) (t : ∀ j, α j → Set (β j)) : (s.pi fun j => t j (update f i a j)) = s.pi fun j => t j (f j) := (pi_congr rfl) fun j hj => by - rw [update_noteq] + rw [update_of_ne] exact fun h => hi (h ▸ hj) theorem pi_update_of_mem [DecidableEq ι] (hi : i ∈ s) (f : ∀ j, α j) (a : α i) @@ -760,7 +760,7 @@ theorem pi_update_of_mem [DecidableEq ι] (hi : i ∈ s) (f : ∀ j, α j) (a : (s.pi fun j => t j (update f i a j)) = ({i} ∪ s \ {i}).pi fun j => t j (update f i a j) := by rw [union_diff_self, union_eq_self_of_subset_left (singleton_subset_iff.2 hi)] _ = { x | x i ∈ t i a } ∩ (s \ {i}).pi fun j => t j (f j) := by - rw [union_pi, singleton_pi', update_same, pi_update_of_not_mem]; simp + rw [union_pi, singleton_pi', update_self, pi_update_of_not_mem]; simp theorem univ_pi_update [DecidableEq ι] {β : ι → Type*} (i : ι) (f : ∀ j, α j) (a : α i) (t : ∀ j, α j → Set (β j)) : @@ -780,7 +780,7 @@ theorem eval_image_univ_pi_subset : eval i '' pi univ t ⊆ t i := theorem subset_eval_image_pi (ht : (s.pi t).Nonempty) (i : ι) : t i ⊆ eval i '' s.pi t := by classical obtain ⟨f, hf⟩ := ht - refine fun y hy => ⟨update f i y, fun j hj => ?_, update_same _ _ _⟩ + refine fun y hy => ⟨update f i y, fun j hj => ?_, update_self ..⟩ obtain rfl | hji := eq_or_ne j i <;> simp [*, hf _ hj] theorem eval_image_pi (hs : i ∈ s) (ht : (s.pi t).Nonempty) : eval i '' s.pi t = t i := @@ -859,7 +859,7 @@ theorem update_preimage_pi [DecidableEq ι] {f : ∀ i, α i} (hi : i ∈ s) simp · obtain rfl | h := eq_or_ne j i · simpa - · rw [update_noteq h] + · rw [update_of_ne h] exact hf j hj h theorem update_image [DecidableEq ι] (x : (i : ι) → β i) (i : ι) (s : Set (β i)) : diff --git a/Mathlib/Data/Sigma/Basic.lean b/Mathlib/Data/Sigma/Basic.lean index 552e884b81259..888d19b2ed8d8 100644 --- a/Mathlib/Data/Sigma/Basic.lean +++ b/Mathlib/Data/Sigma/Basic.lean @@ -169,11 +169,11 @@ theorem Sigma.curry_update {γ : ∀ a, β a → Type*} [DecidableEq α] [∀ a, obtain rfl | ha := eq_or_ne ia ja · obtain rfl | hb := eq_or_ne ib jb · simp - · simp only [update_same] - rw [Function.update_noteq (mt _ hb.symm), Function.update_noteq hb.symm] + · simp only [update_self] + rw [Function.update_of_ne (mt _ hb.symm), Function.update_of_ne hb.symm] rintro h injection h - · rw [Function.update_noteq (ne_of_apply_ne Sigma.fst _), Function.update_noteq] + · rw [Function.update_of_ne (ne_of_apply_ne Sigma.fst _), Function.update_of_ne] · exact ha.symm · exact ha.symm diff --git a/Mathlib/Data/Sum/Basic.lean b/Mathlib/Data/Sum/Basic.lean index 6cf7dd465f694..9631344ff2cbf 100644 --- a/Mathlib/Data/Sum/Basic.lean +++ b/Mathlib/Data/Sum/Basic.lean @@ -88,7 +88,7 @@ theorem update_inl_comp_inr [DecidableEq (α ⊕ β)] {f : α ⊕ β → γ} {i theorem update_inl_apply_inr [DecidableEq (α ⊕ β)] {f : α ⊕ β → γ} {i : α} {j : β} {x : γ} : update f (inl i) x (inr j) = f (inr j) := - Function.update_noteq inr_ne_inl _ _ + Function.update_of_ne inr_ne_inl .. @[simp] theorem update_inr_comp_inl [DecidableEq (α ⊕ β)] {f : α ⊕ β → γ} {i : β} {x : γ} : @@ -97,7 +97,7 @@ theorem update_inr_comp_inl [DecidableEq (α ⊕ β)] {f : α ⊕ β → γ} {i theorem update_inr_apply_inl [DecidableEq (α ⊕ β)] {f : α ⊕ β → γ} {i : α} {j : β} {x : γ} : update f (inr j) x (inl i) = f (inl i) := - Function.update_noteq inl_ne_inr _ _ + Function.update_of_ne inl_ne_inr .. @[simp] theorem update_inr_comp_inr [DecidableEq β] [DecidableEq (α ⊕ β)] {f : α ⊕ β → γ} {i : β} diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index 3bea4b00685c4..eb4aec7e8dbe9 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -377,24 +377,24 @@ instance Icc_smoothManifoldWithCorners (x y : ℝ) [Fact (x < y)] : · -- `e = left chart`, `e' = right chart` apply M.contDiffOn.congr rintro _ ⟨⟨hz₁, hz₂⟩, ⟨⟨z, hz₀⟩, rfl⟩⟩ - simp only [modelWithCornersEuclideanHalfSpace, IccLeftChart, IccRightChart, update_same, + simp only [modelWithCornersEuclideanHalfSpace, IccLeftChart, IccRightChart, update_self, max_eq_left, hz₀, lt_sub_iff_add_lt, mfld_simps] at hz₁ hz₂ rw [min_eq_left hz₁.le, lt_add_iff_pos_left] at hz₂ ext i rw [Subsingleton.elim i 0] simp only [modelWithCornersEuclideanHalfSpace, IccLeftChart, IccRightChart, *, PiLp.add_apply, - PiLp.neg_apply, max_eq_left, min_eq_left hz₁.le, update_same, mfld_simps] + PiLp.neg_apply, max_eq_left, min_eq_left hz₁.le, update_self, mfld_simps] abel · -- `e = right chart`, `e' = left chart` apply M.contDiffOn.congr rintro _ ⟨⟨hz₁, hz₂⟩, ⟨z, hz₀⟩, rfl⟩ simp only [modelWithCornersEuclideanHalfSpace, IccLeftChart, IccRightChart, max_lt_iff, - update_same, max_eq_left hz₀, mfld_simps] at hz₁ hz₂ + update_self, max_eq_left hz₀, mfld_simps] at hz₁ hz₂ rw [lt_sub_comm] at hz₁ ext i rw [Subsingleton.elim i 0] simp only [modelWithCornersEuclideanHalfSpace, IccLeftChart, IccRightChart, PiLp.add_apply, - PiLp.neg_apply, update_same, max_eq_left, hz₀, hz₁.le, mfld_simps] + PiLp.neg_apply, update_self, max_eq_left, hz₀, hz₁.le, mfld_simps] abel ·-- `e = right chart`, `e' = right chart` exact (mem_groupoid_of_pregroupoid.mpr (symm_trans_mem_contDiffGroupoid _)).1 diff --git a/Mathlib/GroupTheory/Complement.lean b/Mathlib/GroupTheory/Complement.lean index 781fb00800227..9bad485b2ce2a 100644 --- a/Mathlib/GroupTheory/Complement.lean +++ b/Mathlib/GroupTheory/Complement.lean @@ -258,10 +258,10 @@ lemma exists_left_transversal (H : Subgroup G) (g : G) : classical refine ⟨Set.range (Function.update Quotient.out _ g), range_mem_leftTransversals fun q => ?_, - Quotient.mk'' g, Function.update_same (Quotient.mk'' g) g Quotient.out⟩ + Quotient.mk'' g, Function.update_self (Quotient.mk'' g) g Quotient.out⟩ by_cases hq : q = Quotient.mk'' g - · exact hq.symm ▸ congr_arg _ (Function.update_same (Quotient.mk'' g) g Quotient.out) - · refine (Function.update_noteq ?_ g Quotient.out) ▸ q.out_eq' + · exact hq.symm ▸ congr_arg _ (Function.update_self (Quotient.mk'' g) g Quotient.out) + · refine (Function.update_of_ne ?_ g Quotient.out) ▸ q.out_eq' exact hq @[to_additive] @@ -270,10 +270,10 @@ lemma exists_right_transversal (H : Subgroup G) (g : G) : classical refine ⟨Set.range (Function.update Quotient.out _ g), range_mem_rightTransversals fun q => ?_, - Quotient.mk'' g, Function.update_same (Quotient.mk'' g) g Quotient.out⟩ + Quotient.mk'' g, Function.update_self (Quotient.mk'' g) g Quotient.out⟩ by_cases hq : q = Quotient.mk'' g - · exact hq.symm ▸ congr_arg _ (Function.update_same (Quotient.mk'' g) g Quotient.out) - · exact Eq.trans (congr_arg _ (Function.update_noteq hq g Quotient.out)) q.out_eq' + · exact hq.symm ▸ congr_arg _ (Function.update_self (Quotient.mk'' g) g Quotient.out) + · exact Eq.trans (congr_arg _ (Function.update_of_ne hq g Quotient.out)) q.out_eq' /-- Given two subgroups `H' ⊆ H`, there exists a left transversal to `H'` inside `H`. -/ @[to_additive "Given two subgroups `H' ⊆ H`, there exists a transversal to `H'` inside `H`"] diff --git a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean index 25dfb8d294ad8..8cd6ea043aebc 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean @@ -930,8 +930,8 @@ theorem affineCombination_mem_affineSpan [Nontrivial k] {s : Finset ι} {w : ι simp only [w1, Function.const_zero, Finset.sum_update_of_mem hi1, Pi.zero_apply, Finset.sum_const_zero, add_zero] have hw1s : s.affineCombination k p w1 = p i1 := - s.affineCombination_of_eq_one_of_eq_zero w1 p hi1 (Function.update_same _ _ _) fun _ _ hne => - Function.update_noteq hne _ _ + s.affineCombination_of_eq_one_of_eq_zero w1 p hi1 (Function.update_self ..) fun _ _ hne => + Function.update_of_ne hne .. have hv : s.affineCombination k p w -ᵥ p i1 ∈ (affineSpan k (Set.range p)).direction := by rw [direction_affineSpan, ← hw1s, Finset.affineCombination_vsub] apply weightedVSub_mem_vectorSpan @@ -1014,7 +1014,7 @@ theorem eq_affineCombination_of_mem_affineSpan {p1 : P} {p : ι → P} add_zero] have hw0s : s'.affineCombination k p w0 = p i0 := s'.affineCombination_of_eq_one_of_eq_zero w0 p (Finset.mem_insert_self _ _) - (Function.update_same _ _ _) fun _ _ hne => Function.update_noteq hne _ _ + (Function.update_self ..) fun _ _ hne => Function.update_of_ne hne _ _ refine ⟨s', w0 + w', ?_, ?_⟩ · simp [Pi.add_apply, Finset.sum_add_distrib, hw0, h'] · rw [add_comm, ← Finset.weightedVSub_vadd_affineCombination, hw0s, hs', vsub_vadd] diff --git a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean index 1adcc9da63cc8..aa46db426e4f3 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean @@ -220,8 +220,8 @@ theorem affineIndependent_iff_indicator_eq_of_affineCombination_eq (p : ι → P rw [Finset.sum_update_of_mem hi0] simp only [Finset.sum_const_zero, add_zero, const_apply] have hw1s : s.affineCombination k p w1 = p i0 := - s.affineCombination_of_eq_one_of_eq_zero w1 p hi0 (Function.update_same _ _ _) - fun _ _ hne => Function.update_noteq hne _ _ + s.affineCombination_of_eq_one_of_eq_zero w1 p hi0 (Function.update_self ..) + fun _ _ hne => Function.update_of_ne hne .. let w2 := w + w1 have hw2 : ∑ i ∈ s, w2 i = 1 := by simp_all only [w2, Pi.add_apply, Finset.sum_add_distrib, zero_add] diff --git a/Mathlib/LinearAlgebra/Alternating/Basic.lean b/Mathlib/LinearAlgebra/Alternating/Basic.lean index 438d06668b433..cde969096e601 100644 --- a/Mathlib/LinearAlgebra/Alternating/Basic.lean +++ b/Mathlib/LinearAlgebra/Alternating/Basic.lean @@ -613,12 +613,12 @@ Various properties of reordered and repeated inputs which follow from theorem map_update_self [DecidableEq ι] {i j : ι} (hij : i ≠ j) : f (Function.update v i (v j)) = 0 := - f.map_eq_zero_of_eq _ (by rw [Function.update_same, Function.update_noteq hij.symm]) hij + f.map_eq_zero_of_eq _ (by rw [Function.update_self, Function.update_of_ne hij.symm]) hij theorem map_update_update [DecidableEq ι] {i j : ι} (hij : i ≠ j) (m : M) : f (Function.update (Function.update v i m) j m) = 0 := f.map_eq_zero_of_eq _ - (by rw [Function.update_same, Function.update_noteq hij, Function.update_same]) hij + (by rw [Function.update_self, Function.update_of_ne hij, Function.update_self]) hij theorem map_swap_add [DecidableEq ι] {i j : ι} (hij : i ≠ j) : f (v ∘ Equiv.swap i j) + f v = 0 := by diff --git a/Mathlib/LinearAlgebra/Matrix/Basis.lean b/Mathlib/LinearAlgebra/Matrix/Basis.lean index f7bd16c3469f5..52990f053c467 100644 --- a/Mathlib/LinearAlgebra/Matrix/Basis.lean +++ b/Mathlib/LinearAlgebra/Matrix/Basis.lean @@ -80,8 +80,8 @@ theorem toMatrix_update [DecidableEq ι'] (x : M) : ext i' k rw [Basis.toMatrix, Matrix.updateCol_apply, e.toMatrix_apply] split_ifs with h - · rw [h, update_same j x v] - · rw [update_noteq h] + · rw [h, update_self j x v] + · rw [update_of_ne h] /-- The basis constructed by `unitsSMul` has vectors given by a diagonal matrix. -/ @[simp] diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index d12337d2b24a3..5de11b6d8fc35 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -167,7 +167,7 @@ theorem map_coord_zero {m : ∀ i, M₁ i} (i : ι) (h : m i = 0) : f m = 0 := b @[simp] theorem map_update_zero [DecidableEq ι] (m : ∀ i, M₁ i) (i : ι) : f (update m i 0) = 0 := - f.map_coord_zero i (update_same i 0 m) + f.map_coord_zero i (update_self i 0 m) @[simp] theorem map_zero [Nonempty ι] : f 0 = 0 := by @@ -500,14 +500,14 @@ theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, intro i by_cases hi : i = i₀ · rw [hi] - simp only [B, sdiff_subset, update_same] - · simp only [B, hi, update_noteq, Ne, not_false_iff, Finset.Subset.refl] + simp only [B, sdiff_subset, update_self] + · simp only [B, hi, update_of_ne, Ne, not_false_iff, Finset.Subset.refl] have C_subset_A : ∀ i, C i ⊆ A i := by intro i by_cases hi : i = i₀ · rw [hi] - simp only [C, hj₂, Finset.singleton_subset_iff, update_same] - · simp only [C, hi, update_noteq, Ne, not_false_iff, Finset.Subset.refl] + simp only [C, hj₂, Finset.singleton_subset_iff, update_self] + · simp only [C, hi, update_of_ne, Ne, not_false_iff, Finset.Subset.refl] -- split the sum at `i₀` as the sum over `B i₀` plus the sum over `C i₀`, to use additivity. have A_eq_BC : (fun i => ∑ j ∈ A i, g i j) = @@ -515,9 +515,9 @@ theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, ((∑ j ∈ B i₀, g i₀ j) + ∑ j ∈ C i₀, g i₀ j) := by ext i by_cases hi : i = i₀ - · rw [hi, update_same] + · rw [hi, update_self] have : A i₀ = B i₀ ∪ C i₀ := by - simp only [B, C, Function.update_same, Finset.sdiff_union_self_eq_union] + simp only [B, C, Function.update_self, Finset.sdiff_union_self_eq_union] symm simp only [hj₂, Finset.singleton_subset_iff, Finset.union_eq_left] rw [this] @@ -526,7 +526,7 @@ theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, simpa [C] using hj rw [this] simp only [B, mem_sdiff, eq_self_iff_true, not_true, not_false_iff, Finset.mem_singleton, - update_same, and_false] + update_self, and_false] · simp [hi] have Beq : Function.update (fun i => ∑ j ∈ A i, g i j) i₀ (∑ j ∈ B i₀, g i₀ j) = fun i => @@ -534,22 +534,22 @@ theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, ext i by_cases hi : i = i₀ · rw [hi] - simp only [update_same] - · simp only [B, hi, update_noteq, Ne, not_false_iff] + simp only [update_self] + · simp only [B, hi, update_of_ne, Ne, not_false_iff] have Ceq : Function.update (fun i => ∑ j ∈ A i, g i j) i₀ (∑ j ∈ C i₀, g i₀ j) = fun i => ∑ j ∈ C i, g i j := by ext i by_cases hi : i = i₀ · rw [hi] - simp only [update_same] - · simp only [C, hi, update_noteq, Ne, not_false_iff] + simp only [update_self] + · simp only [C, hi, update_of_ne, Ne, not_false_iff] -- Express the inductive assumption for `B` have Brec : (f fun i => ∑ j ∈ B i, g i j) = ∑ r ∈ piFinset B, f fun i => g i (r i) := by have : ∑ i, #(B i) < ∑ i, #(A i) := by refine sum_lt_sum (fun i _ => card_le_card (B_subset_A i)) ⟨i₀, mem_univ _, ?_⟩ have : {j₂} ⊆ A i₀ := by simp [hj₂] - simp only [B, Finset.card_sdiff this, Function.update_same, Finset.card_singleton] + simp only [B, Finset.card_sdiff this, Function.update_self, Finset.card_singleton] exact Nat.pred_lt (ne_of_gt (lt_trans Nat.zero_lt_one hi₀)) rw [h] at this exact IH _ this B rfl @@ -708,14 +708,14 @@ lemma domDomRestrict_aux {ι} [DecidableEq ι] (P : ι → Prop) [DecidablePred Function.update (fun j => if h : P j then x ⟨j, h⟩ else z ⟨j, h⟩) i c := by ext j by_cases h : j = i - · rw [h, Function.update_same] - simp only [i.2, update_same, dite_true] - · rw [Function.update_noteq h] + · rw [h, Function.update_self] + simp only [i.2, update_self, dite_true] + · rw [Function.update_of_ne h] by_cases h' : P j · simp only [h', ne_eq, Subtype.mk.injEq, dite_true] have h'' : ¬ ⟨j, h'⟩ = i := fun he => by apply_fun (fun x => x.1) at he; exact h he - rw [Function.update_noteq h''] + rw [Function.update_of_ne h''] · simp only [h', ne_eq, Subtype.mk.injEq, dite_false] lemma domDomRestrict_aux_right {ι} [DecidableEq ι] (P : ι → Prop) [DecidablePred P] {M₁ : ι → Type*} @@ -1002,8 +1002,8 @@ lemma iteratedFDeriv_aux {ι} {M₁ : ι → Type*} {α : Type*} [DecidableEq α (fun i ↦ update (fun j ↦ m (e.symm j) j) (e a) (z (e a)) i) := by ext i rcases eq_or_ne a (e.symm i) with rfl | hne - · rw [Equiv.apply_symm_apply e i, update_same, update_same] - · rw [update_noteq hne.symm, update_noteq fun h ↦ (Equiv.symm_apply_apply .. ▸ h ▸ hne) rfl] + · rw [Equiv.apply_symm_apply e i, update_self, update_self] + · rw [update_of_ne hne.symm, update_of_ne fun h ↦ (Equiv.symm_apply_apply .. ▸ h ▸ hne) rfl] /-- One of the components of the iterated derivative of a multilinear map. Given a bijection `e` between a type `α` (typically `Fin k`) and a subset `s` of `ι`, this component is a multilinear map @@ -1277,7 +1277,7 @@ variable [Semiring R] [∀ i, AddCommGroup (M₁ i)] [AddCommGroup M₂] [∀ i, theorem map_update_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_update_add, neg_add_cancel, f.map_coord_zero i (update_same i 0 m)] + rw [← MultilinearMap.map_update_add, neg_add_cancel, f.map_coord_zero i (update_self i 0 m)] @[deprecated (since := "2024-11-03")] protected alias map_neg := MultilinearMap.map_update_neg @@ -1304,13 +1304,13 @@ lemma map_sub_map_piecewise [LinearOrder ι] (a b : (i : ι) → M₁ i) (s : Fi simp_rw [s.mem_insert] congr 1 · congr; ext i; split_ifs with h₁ h₂ - · rw [update_noteq, Finset.piecewise_eq_of_not_mem] + · rw [update_of_ne, Finset.piecewise_eq_of_not_mem] · exact fun h ↦ (hk i h).not_lt (h₁ <| .inr h) · exact fun h ↦ (h₁ <| .inl h).ne h · cases h₂ - rw [update_same, s.piecewise_eq_of_not_mem _ _ (lt_irrefl _ <| hk k ·)] + rw [update_self, s.piecewise_eq_of_not_mem _ _ (lt_irrefl _ <| hk k ·)] · push_neg at h₁ - rw [update_noteq (Ne.symm h₂), s.piecewise_eq_of_mem _ _ (h₁.1.resolve_left <| Ne.symm h₂)] + rw [update_of_ne (Ne.symm h₂), s.piecewise_eq_of_mem _ _ (h₁.1.resolve_left <| Ne.symm h₂)] · apply sum_congr rfl; intro i hi; congr; ext j; congr 1; apply propext simp_rw [imp_iff_not_or, not_or]; apply or_congr_left' intro h; rw [and_iff_right]; rintro rfl; exact h (hk i hi) @@ -1414,8 +1414,8 @@ def LinearMap.uncurryLeft (f : M 0 →ₗ[R] MultilinearMap R (fun i : Fin n => rw [Subsingleton.elim dec (by clear dec; infer_instance)]; clear dec by_cases h : i = 0 · subst i - simp only [update_same, map_add, tail_update_zero, MultilinearMap.add_apply] - · simp_rw [update_noteq (Ne.symm h)] + simp only [update_self, map_add, tail_update_zero, MultilinearMap.add_apply] + · simp_rw [update_of_ne (Ne.symm h)] revert x y rw [← succ_pred i h] intro x y @@ -1425,8 +1425,8 @@ def LinearMap.uncurryLeft (f : M 0 →ₗ[R] MultilinearMap R (fun i : Fin n => rw [Subsingleton.elim dec (by clear dec; infer_instance)]; clear dec by_cases h : i = 0 · subst i - simp only [update_same, map_smul, tail_update_zero, MultilinearMap.smul_apply] - · simp_rw [update_noteq (Ne.symm h)] + simp only [update_self, map_smul, tail_update_zero, MultilinearMap.smul_apply] + · simp_rw [update_of_ne (Ne.symm h)] revert x rw [← succ_pred i h] intro x @@ -1511,7 +1511,7 @@ def MultilinearMap.uncurryRight rw [Subsingleton.elim dec (by clear dec; infer_instance)]; clear dec by_cases h : i.val < n · have : last n ≠ i := Ne.symm (ne_of_lt h) - simp_rw [update_noteq this] + simp_rw [update_of_ne this] revert x y rw [(castSucc_castLT i h).symm] intro x y @@ -1520,13 +1520,13 @@ def MultilinearMap.uncurryRight · revert x y rw [eq_last_of_not_lt h] intro x y - simp_rw [init_update_last, update_same, LinearMap.map_add] + simp_rw [init_update_last, update_self, LinearMap.map_add] map_update_smul' {dec} m i c x := by -- Porting note: `clear` not necessary in Lean 3 due to not being in the instance cache rw [Subsingleton.elim dec (by clear dec; infer_instance)]; clear dec by_cases h : i.val < n · have : last n ≠ i := Ne.symm (ne_of_lt h) - simp_rw [update_noteq this] + simp_rw [update_of_ne this] revert x rw [(castSucc_castLT i h).symm] intro x @@ -1535,7 +1535,7 @@ def MultilinearMap.uncurryRight · revert x rw [eq_last_of_not_lt h] intro x - simp_rw [update_same, init_update_last, map_smul] + simp_rw [update_self, init_update_last, map_smul] @[simp] theorem MultilinearMap.uncurryRight_apply @@ -1768,9 +1768,9 @@ def map [Nonempty ι] (f : MultilinearMap R M₁ M₂) (p : ∀ i, Submodule R ( let ⟨i⟩ := ‹Nonempty ι› letI := Classical.decEq ι refine ⟨update x i (c • x i), fun j => if hij : j = i then ?_ else ?_, hf ▸ ?_⟩ - · rw [hij, update_same] + · rw [hij, update_self] exact (p i).smul_mem _ (hx i) - · rw [update_noteq hij] + · rw [update_of_ne hij] exact hx j · rw [f.map_update_smul, update_eq_self] diff --git a/Mathlib/LinearAlgebra/Pi.lean b/Mathlib/LinearAlgebra/Pi.lean index 079b058aab3d1..4c1ff2962f310 100644 --- a/Mathlib/LinearAlgebra/Pi.lean +++ b/Mathlib/LinearAlgebra/Pi.lean @@ -288,8 +288,8 @@ def diag (i j : ι) : φ i →ₗ[R] φ j := theorem update_apply (f : (i : ι) → M₂ →ₗ[R] φ i) (c : M₂) (i j : ι) (b : M₂ →ₗ[R] φ i) : (update f i b j) c = update (fun i => f i c) i (b c) j := by by_cases h : j = i - · rw [h, update_same, update_same] - · rw [update_noteq h, update_noteq h] + · rw [h, update_self, update_self] + · rw [update_of_ne h, update_of_ne h] variable (R φ) diff --git a/Mathlib/LinearAlgebra/PiTensorProduct.lean b/Mathlib/LinearAlgebra/PiTensorProduct.lean index 5c38f40c58fc1..50dbf566632b6 100644 --- a/Mathlib/LinearAlgebra/PiTensorProduct.lean +++ b/Mathlib/LinearAlgebra/PiTensorProduct.lean @@ -279,7 +279,7 @@ def tprod : MultilinearMap R s (⨂[R] i, s i) where toFun := tprodCoeff R 1 map_update_add' {_ f} i x y := (add_tprodCoeff (1 : R) f i x y).symm map_update_smul' {_ f} i r x := by - rw [smul_tprodCoeff', ← smul_tprodCoeff (1 : R) _ i, update_idem, update_same] + rw [smul_tprodCoeff', ← smul_tprodCoeff (1 : R) _ i, update_idem, update_self] variable {R} @@ -831,7 +831,7 @@ def subsingletonEquiv [Subsingleton ι] (i₀ : ι) : (⨂[R] _ : ι, M) ≃ₗ[ dsimp only have : ∀ (f : ι → M) (z : M), (fun _ : ι ↦ z) = update f i₀ z := fun f z ↦ by ext i - rw [Subsingleton.elim i i₀, Function.update_same] + rw [Subsingleton.elim i i₀, Function.update_self] refine x.induction_on ?_ ?_ · intro r f simp only [LinearMap.map_smul, LinearMap.id_apply, lift.tprod, ofSubsingleton_apply_apply, diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 093505a183bd2..52c2e153e76e9 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1818,9 +1818,8 @@ theorem piCongrLeft'_update [DecidableEq α] [DecidableEq β] (P : α → Sort*) ext b' rcases eq_or_ne b' b with (rfl | h) · simp - · simp only [Equiv.piCongrLeft'_apply, ne_eq, h, not_false_iff, update_noteq] - rw [update_noteq _] - rw [ne_eq] + · simp only [Equiv.piCongrLeft'_apply, ne_eq, h, not_false_iff, update_of_ne] + rw [update_of_ne] intro h' /- an example of something that should work, or also putting `EmbeddingLike.apply_eq_iff_eq` in the `simp` should too: diff --git a/Mathlib/Logic/Function/Basic.lean b/Mathlib/Logic/Function/Basic.lean index a1af4bbb91992..a59d0b66dcbae 100644 --- a/Mathlib/Logic/Function/Basic.lean +++ b/Mathlib/Logic/Function/Basic.lean @@ -489,13 +489,17 @@ def update (f : ∀ a, β a) (a' : α) (v : β a') (a : α) : β a := if h : a = a' then Eq.ndrec v h.symm else f a @[simp] -theorem update_same (a : α) (v : β a) (f : ∀ a, β a) : update f a v a = v := +theorem update_self (a : α) (v : β a) (f : ∀ a, β a) : update f a v a = v := dif_pos rfl +@[deprecated (since := "2024-12-28")] alias update_same := update_self + @[simp] -theorem update_noteq {a a' : α} (h : a ≠ a') (v : β a') (f : ∀ a, β a) : update f a' v a = f a := +theorem update_of_ne {a a' : α} (h : a ≠ a') (v : β a') (f : ∀ a, β a) : update f a' v a = f a := dif_neg h +@[deprecated (since := "2024-12-28")] alias update_noteq := update_of_ne + /-- On non-dependent functions, `Function.update` can be expressed as an `ite` -/ theorem update_apply {β : Sort*} (f : α → β) (a' : α) (b : β) (a : α) : update f a' b a = if a = a' then b else f a := by @@ -504,20 +508,20 @@ theorem update_apply {β : Sort*} (f : α → β) (a' : α) (b : β) (a : α) : @[nontriviality] theorem update_eq_const_of_subsingleton [Subsingleton α] (a : α) (v : α') (f : α → α') : update f a v = const α v := - funext fun a' ↦ Subsingleton.elim a a' ▸ update_same _ _ _ + funext fun a' ↦ Subsingleton.elim a a' ▸ update_self .. theorem surjective_eval {α : Sort u} {β : α → Sort v} [h : ∀ a, Nonempty (β a)] (a : α) : Surjective (eval a : (∀ a, β a) → β a) := fun b ↦ ⟨@update _ _ (Classical.decEq α) (fun a ↦ (h a).some) a b, - @update_same _ _ (Classical.decEq α) _ _ _⟩ + @update_self _ _ (Classical.decEq α) _ _ _⟩ theorem update_injective (f : ∀ a, β a) (a' : α) : Injective (update f a') := fun v v' h ↦ by have := congr_fun h a' - rwa [update_same, update_same] at this + rwa [update_self, update_self] at this lemma forall_update_iff (f : ∀a, β a) {a : α} {b : β a} (p : ∀a, β a → Prop) : (∀ x, p x (update f a b x)) ↔ p a b ∧ ∀ x, x ≠ a → p x (f x) := by - rw [← and_forall_ne a, update_same] + rw [← and_forall_ne a, update_self] simp +contextual theorem exists_update_iff (f : ∀ a, β a) {a : α} {b : β a} (p : ∀ a, β a → Prop) : @@ -547,7 +551,7 @@ theorem update_eq_self (a : α) (f : ∀ a, β a) : update f a (f a) = f := theorem update_comp_eq_of_forall_ne' {α'} (g : ∀ a, β a) {f : α' → α} {i : α} (a : β i) (h : ∀ x, f x ≠ i) : (fun j ↦ (update g i a) (f j)) = fun j ↦ g (f j) := - funext fun _ ↦ update_noteq (h _) _ _ + funext fun _ ↦ update_of_ne (h _) _ _ variable [DecidableEq α'] @@ -558,7 +562,7 @@ theorem update_comp_eq_of_forall_ne {α β : Sort*} (g : α' → β) {f : α → theorem update_comp_eq_of_injective' (g : ∀ a, β a) {f : α' → α} (hf : Function.Injective f) (i : α') (a : β (f i)) : (fun j ↦ update g (f i) a (f j)) = update (fun i ↦ g (f i)) i a := - eq_update_iff.2 ⟨update_same _ _ _, fun _ hj ↦ update_noteq (hf.ne hj) _ _⟩ + eq_update_iff.2 ⟨update_self .., fun _ hj ↦ update_of_ne (hf.ne hj) _ _⟩ /-- Non-dependent version of `Function.update_comp_eq_of_injective'` -/ theorem update_comp_eq_of_injective {β : Sort*} (g : α' → β) {f : α → α'} diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index fb04b17543a7c..963c8dcd5082f 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -118,8 +118,8 @@ theorem generateFrom_pi_eq {C : ∀ i, Set (Set (α i))} (hC : ∀ i, IsCountabl intro n; apply measurableSet_generateFrom apply mem_image_of_mem; intro j _; dsimp only by_cases h : j = i - · subst h; rwa [update_same] - · rw [update_noteq h]; apply h1t + · subst h; rwa [update_self] + · rw [update_of_ne h]; apply h1t · apply generateFrom_le; rintro _ ⟨s, hs, rfl⟩ rw [univ_pi_eq_iInter]; apply MeasurableSet.iInter; intro i apply @measurable_pi_apply _ _ (fun i => generateFrom (C i)) diff --git a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean index dbb7952724845..7dbef7d066822 100644 --- a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean +++ b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean @@ -1152,14 +1152,14 @@ theorem integral_eq_sub_of_hasDerivAt_of_tendsto (hab : a < b) {fa fb} refine fun x hx => (hderiv x hx).congr_of_eventuallyEq ?_ filter_upwards [Ioo_mem_nhds hx.1 hx.2] with _ hy unfold F - rw [update_noteq hy.2.ne, update_noteq hy.1.ne'] + rw [update_of_ne hy.2.ne, update_of_ne hy.1.ne'] have hcont : ContinuousOn F (Icc a b) := by rw [continuousOn_update_iff, continuousOn_update_iff, Icc_diff_right, Ico_diff_left] refine ⟨⟨fun z hz => (hderiv z hz).continuousAt.continuousWithinAt, ?_⟩, ?_⟩ · exact fun _ => ha.mono_left (nhdsWithin_mono _ Ioo_subset_Ioi_self) · rintro - refine (hb.congr' ?_).mono_left (nhdsWithin_mono _ Ico_subset_Iio_self) - filter_upwards [Ioo_mem_nhdsLT hab] with _ hz using (update_noteq hz.1.ne' _ _).symm + filter_upwards [Ioo_mem_nhdsLT hab] with _ hz using (update_of_ne hz.1.ne' _ _).symm simpa [F, hab.ne, hab.ne'] using integral_eq_sub_of_hasDerivAt_of_le hab.le hcont Fderiv hint /-- Fundamental theorem of calculus-2: If `f : ℝ → E` is differentiable at every `x` in `[a, b]` and diff --git a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean index fa376363a2403..99c22b52d22c3 100644 --- a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean +++ b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean @@ -1265,11 +1265,11 @@ theorem integral_Ioi_deriv_mul_eq_sub intro x (hx : a < x) apply ((hu x hx).mul (hv x hx)).congr_of_eventuallyEq filter_upwards [eventually_ne_nhds hx.ne.symm] with y hy - exact Function.update_noteq hy a' (u * v) + exact Function.update_of_ne hy a' (u * v) have htendsto : Tendsto f atTop (𝓝 b') := by apply h_infty.congr' filter_upwards [eventually_ne_atTop a] with x hx - exact (Function.update_noteq hx a' (u * v)).symm + exact (Function.update_of_ne hx a' (u * v)).symm simpa using integral_Ioi_of_hasDerivAt_of_tendsto (continuousWithinAt_update_same.mpr h_zero) hderiv huv htendsto @@ -1296,11 +1296,11 @@ theorem integral_Iic_deriv_mul_eq_sub intro x hx apply ((hu x hx).mul (hv x hx)).congr_of_eventuallyEq filter_upwards [Iio_mem_nhds hx] with x (hx : x < a) - exact Function.update_noteq (ne_of_lt hx) a' (u * v) + exact Function.update_of_ne (ne_of_lt hx) a' (u * v) have htendsto : Tendsto f atBot (𝓝 b') := by apply h_infty.congr' filter_upwards [Iio_mem_atBot a] with x (hx : x < a) - exact (Function.update_noteq (ne_of_lt hx) a' (u * v)).symm + exact (Function.update_of_ne (ne_of_lt hx) a' (u * v)).symm simpa using integral_Iic_of_hasDerivAt_of_tendsto (continuousWithinAt_update_same.mpr h_zero) hderiv huv htendsto diff --git a/Mathlib/MeasureTheory/Integral/Marginal.lean b/Mathlib/MeasureTheory/Integral/Marginal.lean index 2c027ef3bc5e7..483452c38db13 100644 --- a/Mathlib/MeasureTheory/Integral/Marginal.lean +++ b/Mathlib/MeasureTheory/Integral/Marginal.lean @@ -112,7 +112,7 @@ theorem lmarginal_update_of_mem {i : δ} (hi : i ∈ s) apply lmarginal_congr intro j hj have : j ≠ i := by rintro rfl; exact hj hi - apply update_noteq this + apply update_of_ne this variable {μ} in theorem lmarginal_singleton (f : (∀ i, π i) → ℝ≥0∞) (i : δ) : diff --git a/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean b/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean index 033ad5dde6fc6..4b59c880997f0 100644 --- a/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean +++ b/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean @@ -334,7 +334,7 @@ noncomputable abbrev LFunctionTrivChar₁ : ℂ → ℂ := (∏ p ∈ n.primeFactors, (1 - (p : ℂ)⁻¹)) lemma LFunctionTrivChar₁_apply_one_ne_zero : LFunctionTrivChar₁ n 1 ≠ 0 := by - simp only [Function.update_same] + simp only [Function.update_self] refine Finset.prod_ne_zero_iff.mpr fun p hp ↦ ?_ simpa only [ne_eq, sub_ne_zero, one_eq_inv, Nat.cast_eq_one] using (Nat.prime_of_mem_primeFactors hp).ne_one @@ -345,7 +345,7 @@ lemma differentiable_LFunctionTrivChar₁ : Differentiable ℂ (LFunctionTrivCha ← differentiableOn_compl_singleton_and_continuousAt_iff (c := 1) Filter.univ_mem] refine ⟨DifferentiableOn.congr (f := fun s ↦ (s - 1) * LFunctionTrivChar n s) (fun _ hs ↦ DifferentiableAt.differentiableWithinAt <| by fun_prop (disch := simp_all [hs])) - fun _ hs ↦ Function.update_noteq (Set.mem_diff_singleton.mp hs).2 .., + fun _ hs ↦ Function.update_of_ne (Set.mem_diff_singleton.mp hs).2 .., continuousWithinAt_compl_self.mp ?_⟩ simpa only [continuousWithinAt_compl_self, continuousAt_update_same] using LFunctionTrivChar_residue_one @@ -356,7 +356,7 @@ lemma deriv_LFunctionTrivChar₁_apply_of_ne_one {s : ℂ} (hs : s ≠ 1) : have H : deriv (LFunctionTrivChar₁ n) s = deriv (fun w ↦ (w - 1) * LFunctionTrivChar n w) s := by refine eventuallyEq_iff_exists_mem.mpr ?_ |>.deriv_eq - exact ⟨_, isOpen_ne.mem_nhds hs, fun _ hw ↦ Function.update_noteq (Set.mem_setOf.mp hw) ..⟩ + exact ⟨_, isOpen_ne.mem_nhds hs, fun _ hw ↦ Function.update_of_ne (Set.mem_setOf.mp hw) ..⟩ rw [H, deriv_mul (by fun_prop) (differentiableAt_LFunction _ s (.inl hs)), deriv_sub_const, deriv_id'', one_mul, add_comm] @@ -371,7 +371,7 @@ lemma continuousOn_neg_logDeriv_LFunctionTrivChar₁ : h.continuous.continuousOn fun w hw ↦ ?_).neg rcases eq_or_ne w 1 with rfl | hw' · exact LFunctionTrivChar₁_apply_one_ne_zero _ - · rw [LFunctionTrivChar₁, Function.update_noteq hw', mul_ne_zero_iff] + · rw [LFunctionTrivChar₁, Function.update_of_ne hw', mul_ne_zero_iff] exact ⟨sub_ne_zero_of_ne hw', (Set.mem_setOf.mp hw).resolve_left hw'⟩ end trivial diff --git a/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean b/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean index 81dee4c2c5ae9..9782d336636d6 100644 --- a/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean +++ b/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean @@ -586,7 +586,7 @@ lemma differentiableAt_update_of_residue · -- Easy case : `s ≠ 0` refine (claim s hs hs').congr_of_eventuallyEq ?_ filter_upwards [isOpen_compl_singleton.mem_nhds hs] with x hx - simp only [Function.update_noteq hx] + simp only [Function.update_of_ne hx] · -- Hard case : `s = 0` simp_rw [← claim2.limUnder_eq] have S_nhds : {(1 : ℂ)}ᶜ ∈ 𝓝 (0 : ℂ) := isOpen_compl_singleton.mem_nhds hs' @@ -609,14 +609,14 @@ lemma hurwitzZetaEven_def_of_ne_or_ne {a : UnitAddCircle} {s : ℂ} (h : a ≠ 0 hurwitzZetaEven a s = completedHurwitzZetaEven a s / Gammaℝ s := by rw [hurwitzZetaEven] rcases ne_or_eq s 0 with h | rfl - · rw [Function.update_noteq h] - · simpa only [Gammaℝ, Function.update_same, neg_zero, zero_div, cpow_zero, Complex.Gamma_zero, + · rw [Function.update_of_ne h] + · simpa only [Gammaℝ, Function.update_self, neg_zero, zero_div, cpow_zero, Complex.Gamma_zero, mul_zero, div_zero, ite_eq_right_iff, div_eq_zero_iff, neg_eq_zero, one_ne_zero, OfNat.ofNat_ne_zero, or_self, imp_false, ne_eq, not_true_eq_false, or_false] using h lemma hurwitzZetaEven_apply_zero (a : UnitAddCircle) : hurwitzZetaEven a 0 = if a = 0 then -1 / 2 else 0 := - Function.update_same _ _ _ + Function.update_self .. lemma hurwitzZetaEven_neg (a : UnitAddCircle) (s : ℂ) : hurwitzZetaEven (-a) s = hurwitzZetaEven a s := by @@ -627,7 +627,7 @@ theorem hurwitzZetaEven_neg_two_mul_nat_add_one (a : UnitAddCircle) (n : ℕ) : hurwitzZetaEven a (-2 * (n + 1)) = 0 := by have : (-2 : ℂ) * (n + 1) ≠ 0 := mul_ne_zero (neg_ne_zero.mpr two_ne_zero) (Nat.cast_add_one_ne_zero n) - rw [hurwitzZetaEven, Function.update_noteq this, + rw [hurwitzZetaEven, Function.update_of_ne this, Gammaℝ_eq_zero_iff.mpr ⟨n + 1, by rw [neg_mul, Nat.cast_add_one]⟩, div_zero] /-- The Hurwitz zeta function is differentiable everywhere except at `s = 1`. This is true @@ -656,7 +656,7 @@ lemma differentiableAt_hurwitzZetaEven_sub_one_div (a : UnitAddCircle) : (fun s ↦ completedHurwitzZetaEven a s / Gammaℝ s - 1 / (s - 1) / Gammaℝ s) 1 by apply this.congr_of_eventuallyEq filter_upwards [eventually_ne_nhds one_ne_zero] with x hx - rw [hurwitzZetaEven, Function.update_noteq hx] + rw [hurwitzZetaEven, Function.update_of_ne hx] simp_rw [← sub_div, div_eq_mul_inv _ (Gammaℝ _)] refine DifferentiableAt.mul ?_ differentiable_Gammaℝ_inv.differentiableAt simp_rw [completedHurwitzZetaEven_eq, sub_sub, add_assoc] @@ -687,7 +687,7 @@ Formula for `hurwitzZetaEven` as a Dirichlet series in the convergence range, wi -/ lemma hasSum_int_hurwitzZetaEven (a : ℝ) {s : ℂ} (hs : 1 < re s) : HasSum (fun n : ℤ ↦ 1 / (↑|n + a| : ℂ) ^ s / 2) (hurwitzZetaEven a s) := by - rw [hurwitzZetaEven, Function.update_noteq (ne_zero_of_one_lt_re hs)] + rw [hurwitzZetaEven, Function.update_of_ne (ne_zero_of_one_lt_re hs)] have := (hasSum_int_completedHurwitzZetaEven a hs).div_const (Gammaℝ s) exact this.congr_fun fun n ↦ by simp only [div_right_comm _ _ (Gammaℝ _), div_self (Gammaℝ_ne_zero_of_re_pos (zero_lt_one.trans hs))] @@ -721,7 +721,7 @@ noncomputable def cosZeta (a : UnitAddCircle) := Function.update (fun s : ℂ ↦ completedCosZeta a s / Gammaℝ s) 0 (-1 / 2) lemma cosZeta_apply_zero (a : UnitAddCircle) : cosZeta a 0 = -1 / 2 := - Function.update_same _ _ _ + Function.update_self .. lemma cosZeta_neg (a : UnitAddCircle) (s : ℂ) : cosZeta (-a) s = cosZeta a s := by @@ -732,7 +732,7 @@ theorem cosZeta_neg_two_mul_nat_add_one (a : UnitAddCircle) (n : ℕ) : cosZeta a (-2 * (n + 1)) = 0 := by have : (-2 : ℂ) * (n + 1) ≠ 0 := mul_ne_zero (neg_ne_zero.mpr two_ne_zero) (Nat.cast_add_one_ne_zero n) - rw [cosZeta, Function.update_noteq this, + rw [cosZeta, Function.update_of_ne this, Gammaℝ_eq_zero_iff.mpr ⟨n + 1, by rw [neg_mul, Nat.cast_add_one]⟩, div_zero] /-- The cosine zeta function is differentiable everywhere, except at `s = 1` if `a = 0`. -/ @@ -744,7 +744,7 @@ lemma differentiableAt_cosZeta (a : UnitAddCircle) {s : ℂ} (hs' : s ≠ 1 ∨ · apply ((differentiableAt_completedCosZeta a one_ne_zero hs').mul (differentiable_Gammaℝ_inv.differentiableAt)).congr_of_eventuallyEq filter_upwards [isOpen_compl_singleton.mem_nhds one_ne_zero] with x hx - simp_rw [cosZeta, Function.update_noteq hx, div_eq_mul_inv] + simp_rw [cosZeta, Function.update_of_ne hx, div_eq_mul_inv] /-- If `a ≠ 0` then the cosine zeta function is entire. -/ lemma differentiable_cosZeta_of_ne_zero {a : UnitAddCircle} (ha : a ≠ 0) : @@ -754,7 +754,7 @@ lemma differentiable_cosZeta_of_ne_zero {a : UnitAddCircle} (ha : a ≠ 0) : /-- Formula for `cosZeta` as a Dirichlet series in the convergence range, with sum over `ℤ`. -/ lemma hasSum_int_cosZeta (a : ℝ) {s : ℂ} (hs : 1 < re s) : HasSum (fun n : ℤ ↦ cexp (2 * π * I * a * n) / ↑|n| ^ s / 2) (cosZeta a s) := by - rw [cosZeta, Function.update_noteq (ne_zero_of_one_lt_re hs)] + rw [cosZeta, Function.update_of_ne (ne_zero_of_one_lt_re hs)] refine ((hasSum_int_completedCosZeta a hs).div_const (Gammaℝ s)).congr_fun fun n ↦ ?_ rw [mul_div_assoc _ (cexp _), div_right_comm _ (2 : ℂ), mul_div_cancel_left₀ _ (Gammaℝ_ne_zero_of_re_pos (zero_lt_one.trans hs))] @@ -788,7 +788,7 @@ lemma hurwitzZetaEven_one_sub (a : UnitAddCircle) {s : ℂ} rw [hurwitzZetaEven_def_of_ne_or_ne, div_eq_mul_inv] simpa [sub_eq_zero, eq_comm (a := s)] using hs' rw [this, completedHurwitzZetaEven_one_sub, inv_Gammaℝ_one_sub hs, cosZeta, - Function.update_noteq (by simpa using hs 0), ← Gammaℂ] + Function.update_of_ne (by simpa using hs 0), ← Gammaℂ] generalize Gammaℂ s * cos (π * s / 2) = A -- speeds up ring_nf call ring_nf @@ -798,7 +798,7 @@ lemma cosZeta_one_sub (a : UnitAddCircle) {s : ℂ} (hs : ∀ (n : ℕ), s ≠ 1 cosZeta a (1 - s) = 2 * (2 * π) ^ (-s) * Gamma s * cos (π * s / 2) * hurwitzZetaEven a s := by rw [← Gammaℂ] have : cosZeta a (1 - s) = completedCosZeta a (1 - s) * (Gammaℝ (1 - s))⁻¹ := by - rw [cosZeta, Function.update_noteq, div_eq_mul_inv] + rw [cosZeta, Function.update_of_ne, div_eq_mul_inv] simpa [sub_eq_zero] using (hs 0).symm rw [this, completedCosZeta_one_sub, inv_Gammaℝ_one_sub (fun n ↦ by simpa using hs (n + 1)), hurwitzZetaEven_def_of_ne_or_ne (Or.inr (by simpa using hs 1))] diff --git a/Mathlib/NumberTheory/LSeries/Nonvanishing.lean b/Mathlib/NumberTheory/LSeries/Nonvanishing.lean index 0f0d36fcee6b0..24e61f8fb2458 100644 --- a/Mathlib/NumberTheory/LSeries/Nonvanishing.lean +++ b/Mathlib/NumberTheory/LSeries/Nonvanishing.lean @@ -140,14 +140,14 @@ private lemma F_differentiableAt_of_ne (B : BadChar N) {s : ℂ} (hs : s ≠ 1) DifferentiableAt ℂ B.F s := by apply DifferentiableAt.congr_of_eventuallyEq · exact (differentiableAt_riemannZeta hs).mul <| differentiableAt_LFunction B.χ s (.inl hs) - · filter_upwards [eventually_ne_nhds hs] with t ht using Function.update_noteq ht .. + · filter_upwards [eventually_ne_nhds hs] with t ht using Function.update_of_ne ht .. /-- `B.F` agrees with the L-series of `zetaMul χ` on `1 < s.re`. -/ private lemma F_eq_LSeries (B : BadChar N) {s : ℂ} (hs : 1 < s.re) : B.F s = LSeries B.χ.zetaMul s := by rw [F, zetaMul, ← coe_mul, LSeries_convolution'] · have hs' : s ≠ 1 := fun h ↦ by simp only [h, one_re, lt_self_iff_false] at hs - simp only [ne_eq, hs', not_false_eq_true, Function.update_noteq, B.χ.LFunction_eq_LSeries hs] + simp only [ne_eq, hs', not_false_eq_true, Function.update_of_ne, B.χ.LFunction_eq_LSeries hs] congr 1 · simp_rw [← LSeries_zeta_eq_riemannZeta hs, ← natCoe_apply] · exact LSeries_congr s B.χ.apply_eq_toArithmeticFunction_apply @@ -172,8 +172,8 @@ private lemma F_differentiable (B : BadChar N) : Differentiable ℂ B.F := by have : B.F = G * H := by ext1 t rcases eq_or_ne t 1 with rfl | ht - · simp only [F, G, H, Pi.mul_apply, one_mul, Function.update_same] - · simp only [F, G, H, Function.update_noteq ht, mul_comm _ (riemannZeta _), B.hχ, sub_zero, + · simp only [F, G, H, Pi.mul_apply, one_mul, Function.update_self] + · simp only [F, G, H, Function.update_of_ne ht, mul_comm _ (riemannZeta _), B.hχ, sub_zero, Pi.mul_apply, mul_assoc, mul_div_cancel₀ _ (sub_ne_zero.mpr ht)] rw [this] apply ContinuousAt.mul @@ -185,7 +185,7 @@ This is used later to obtain a contradction. -/ private lemma F_neg_two (B : BadChar N) : B.F (-2 : ℝ) = 0 := by have := riemannZeta_neg_two_mul_nat_add_one 0 rw [Nat.cast_zero, zero_add, mul_one] at this - rw [F, ofReal_neg, ofReal_ofNat, Function.update_noteq (mod_cast (by omega : (-2 : ℤ) ≠ 1)), + rw [F, ofReal_neg, ofReal_ofNat, Function.update_of_ne (mod_cast (by omega : (-2 : ℤ) ≠ 1)), this, zero_mul] end BadChar diff --git a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean index 7139b4bb1fc27..a696c88e59f26 100644 --- a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean +++ b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean @@ -363,7 +363,7 @@ lemma eqOn_LFunctionResidueClassAux (ha : IsUnit a) : congrm (?_ + _) have hs₁ : s ≠ 1 := fun h ↦ ((h ▸ hs).trans_eq one_re).false rw [deriv_LFunctionTrivChar₁_apply_of_ne_one _ hs₁, LFunctionTrivChar₁, - Function.update_noteq hs₁, LFunctionTrivChar, add_div, + Function.update_of_ne hs₁, LFunctionTrivChar, add_div, mul_div_mul_left _ _ (sub_ne_zero_of_ne hs₁)] conv_lhs => enter [2, 1]; rw [← mul_one (LFunction ..)] rw [mul_comm _ 1, mul_div_mul_right _ _ <| LFunction_ne_zero_of_one_le_re 1 (.inr hs₁) hs.le] diff --git a/Mathlib/NumberTheory/LSeries/RiemannZeta.lean b/Mathlib/NumberTheory/LSeries/RiemannZeta.lean index 0b29f4b14b30a..8642146695e22 100644 --- a/Mathlib/NumberTheory/LSeries/RiemannZeta.lean +++ b/Mathlib/NumberTheory/LSeries/RiemannZeta.lean @@ -136,11 +136,11 @@ theorem differentiableAt_riemannZeta {s : ℂ} (hs' : s ≠ 1) : DifferentiableA /-- We have `ζ(0) = -1 / 2`. -/ theorem riemannZeta_zero : riemannZeta 0 = -1 / 2 := by - simp_rw [riemannZeta, hurwitzZetaEven, Function.update_same, if_true] + simp_rw [riemannZeta, hurwitzZetaEven, Function.update_self, if_true] lemma riemannZeta_def_of_ne_zero {s : ℂ} (hs : s ≠ 0) : riemannZeta s = completedRiemannZeta s / Gammaℝ s := by - rw [riemannZeta, hurwitzZetaEven, Function.update_noteq hs, completedHurwitzZetaEven_zero] + rw [riemannZeta, hurwitzZetaEven, Function.update_of_ne hs, completedHurwitzZetaEven_zero] /-- The trivial zeroes of the zeta function. -/ theorem riemannZeta_neg_two_mul_nat_add_one (n : ℕ) : riemannZeta (-2 * (n + 1)) = 0 := diff --git a/Mathlib/NumberTheory/LSeries/ZMod.lean b/Mathlib/NumberTheory/LSeries/ZMod.lean index a7b7dd924fbf6..e25eb91d29554 100644 --- a/Mathlib/NumberTheory/LSeries/ZMod.lean +++ b/Mathlib/NumberTheory/LSeries/ZMod.lean @@ -412,7 +412,7 @@ private lemma completedLFunction_one_sub_of_one_lt_even (hΦ : Φ.Even) {s : ℂ simp only [completedLFunction_def_even hΦ, neg_sub, completedHurwitzZetaEven_one_sub, this] -- reduce to equality with un-completed L-functions: suffices ∑ x, Φ x * cosZeta (toAddCircle x) s = LFunction (𝓕 Φ) s by - simpa only [cosZeta, Function.update_noteq hs₀, ← mul_div_assoc, ← sum_div, + simpa only [cosZeta, Function.update_of_ne hs₀, ← mul_div_assoc, ← sum_div, LFunction_eq_completed_div_gammaFactor_even (dft_even_iff.mpr hΦ) _ (.inl hs₀), div_left_inj' (Gammaℝ_ne_zero_of_re_pos (zero_lt_one.trans hs))] -- expand out `LFunction (𝓕 Φ)` and use parity: diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean index 5603d78d81866..0eb57e9a67069 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean @@ -124,9 +124,9 @@ theorem adjust_f {w₁ : InfinitePlace K} (B : ℝ≥0) (hf : ∀ w, w ≠ w₁ ∃ g : InfinitePlace K → ℝ≥0, (∀ w, w ≠ w₁ → g w = f w) ∧ ∏ w, (g w) ^ mult w = B := by let S := ∏ w ∈ Finset.univ.erase w₁, (f w) ^ mult w refine ⟨Function.update f w₁ ((B * S⁻¹) ^ (mult 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)], + · exact fun w hw => Function.update_of_ne hw _ f + · rw [← Finset.mul_prod_erase Finset.univ _ (Finset.mem_univ w₁), Function.update_self, + Finset.prod_congr rfl fun w hw => by rw [Function.update_of_ne (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] · rw [Finset.prod_ne_zero_iff] diff --git a/Mathlib/Order/Atoms.lean b/Mathlib/Order/Atoms.lean index 6d4b89c1b3df3..a9f5023c860eb 100644 --- a/Mathlib/Order/Atoms.lean +++ b/Mathlib/Order/Atoms.lean @@ -1087,11 +1087,11 @@ theorem isAtom_iff {f : ∀ i, π i} [∀ i, PartialOrder (π i)] [∀ i, OrderB have := h (Function.update ⊥ j (f j)) simp only [lt_def, le_def, Pi.eq_bot_iff, and_imp, forall_exists_index] at this simpa using this (fun k => by by_cases h : k = j; { subst h; simp }; simp [h]) i - (by rwa [Function.update_noteq (Ne.symm hj), bot_apply, bot_lt_iff_ne_bot]) j + (by rwa [Function.update_of_ne (Ne.symm hj), bot_apply, bot_lt_iff_ne_bot]) j theorem isAtom_single {i : ι} [DecidableEq ι] [∀ i, PartialOrder (π i)] [∀ i, OrderBot (π i)] {a : π i} (h : IsAtom a) : IsAtom (Function.update (⊥ : ∀ i, π i) i a) := - isAtom_iff.2 ⟨i, by simpa, fun _ hji => Function.update_noteq hji _ _⟩ + isAtom_iff.2 ⟨i, by simpa, fun _ hji => Function.update_of_ne hji ..⟩ theorem isAtom_iff_eq_single [DecidableEq ι] [∀ i, PartialOrder (π i)] [∀ i, OrderBot (π i)] {f : ∀ i, π i} : @@ -1101,7 +1101,7 @@ theorem isAtom_iff_eq_single [DecidableEq ι] [∀ i, PartialOrder (π i)] intro h have ⟨i, h, hbot⟩ := isAtom_iff.1 h refine ⟨_, _, h, funext fun j => if hij : j = i then hij ▸ by simp else ?_⟩ - rw [Function.update_noteq hij, hbot _ hij, bot_apply] + rw [Function.update_of_ne hij, hbot _ hij, bot_apply] case mpr => rintro ⟨i, a, h, rfl⟩ exact isAtom_single h @@ -1135,8 +1135,8 @@ instance isAtomistic [∀ i, CompleteLattice (π i)] [∀ i, IsAtomistic (π i)] case ge => refine sSup_le ?_ rintro _ ⟨⟨_, ⟨j, a, ha, rfl⟩, hle⟩, rfl⟩ - if hij : i = j then ?_ else simp [Function.update_noteq hij] - subst hij; simp only [Function.update_same] + if hij : i = j then ?_ else simp [Function.update_of_ne hij] + subst hij; simp only [Function.update_self] exact le_sSup ⟨ha, by simpa using hle i⟩ instance isCoatomistic [∀ i, CompleteLattice (π i)] [∀ i, IsCoatomistic (π i)] : diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 7649faffca083..2db0d82f5aeed 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -1258,8 +1258,8 @@ instance [∀ i, Preorder (π i)] [∀ i, DenselyOrdered (π i)] : obtain ⟨c, ha, hb⟩ := exists_between hi exact ⟨Function.update a i c, - ⟨le_update_iff.2 ⟨ha.le, fun _ _ ↦ le_rfl⟩, i, by rwa [update_same]⟩, - update_le_iff.2 ⟨hb.le, fun _ _ ↦ hab _⟩, i, by rwa [update_same]⟩⟩ + ⟨le_update_iff.2 ⟨ha.le, fun _ _ ↦ le_rfl⟩, i, by rwa [update_self]⟩, + update_le_iff.2 ⟨hb.le, fun _ _ ↦ hab _⟩, i, by rwa [update_self]⟩⟩ theorem le_of_forall_le_of_dense [LinearOrder α] [DenselyOrdered α] {a₁ a₂ : α} (h : ∀ a, a₂ < a → a₁ ≤ a) : a₁ ≤ a₂ := diff --git a/Mathlib/Order/Bounds/Image.lean b/Mathlib/Order/Bounds/Image.lean index 4ef77e7f9e476..c64d79501a1b0 100644 --- a/Mathlib/Order/Bounds/Image.lean +++ b/Mathlib/Order/Bounds/Image.lean @@ -494,7 +494,7 @@ theorem isLUB_pi {s : Set (∀ a, π a)} {f : ∀ a, π a} : refine ⟨fun H a => ⟨(Function.monotone_eval a).mem_upperBounds_image H.1, fun b hb => ?_⟩, fun H => ⟨?_, ?_⟩⟩ - · suffices h : Function.update f a b ∈ upperBounds s from Function.update_same a b f ▸ H.2 h a + · suffices h : Function.update f a b ∈ upperBounds s from Function.update_self a b f ▸ H.2 h a exact fun g hg => le_update_iff.2 ⟨hb <| mem_image_of_mem _ hg, fun i _ => H.1 hg i⟩ · exact fun g hg a => (H a).1 (mem_image_of_mem _ hg) · exact fun g hg a => (H a).2 ((Function.monotone_eval a).mem_upperBounds_image hg) diff --git a/Mathlib/Order/CompleteLattice/Finset.lean b/Mathlib/Order/CompleteLattice/Finset.lean index b3607813b1f92..e5891f26b9252 100644 --- a/Mathlib/Order/CompleteLattice/Finset.lean +++ b/Mathlib/Order/CompleteLattice/Finset.lean @@ -158,8 +158,8 @@ theorem iInf_finset_image {f : γ → α} {g : α → β} {s : Finset γ} : theorem iSup_insert_update {x : α} {t : Finset α} (f : α → β) {s : β} (hx : x ∉ t) : ⨆ i ∈ insert x t, Function.update f x s i = s ⊔ ⨆ i ∈ t, f i := by - simp only [Finset.iSup_insert, update_same] - rcongr (i hi); apply update_noteq; rintro rfl; exact hx hi + simp only [Finset.iSup_insert, update_self] + rcongr (i hi); apply update_of_ne; rintro rfl; exact hx hi theorem iInf_insert_update {x : α} {t : Finset α} (f : α → β) {s : β} (hx : x ∉ t) : ⨅ i ∈ insert x t, update f x s i = s ⊓ ⨅ i ∈ t, f i := diff --git a/Mathlib/Order/Interval/Set/Pi.lean b/Mathlib/Order/Interval/Set/Pi.lean index 3d63a7437ca29..fb1b5ec01bf9b 100644 --- a/Mathlib/Order/Interval/Set/Pi.lean +++ b/Mathlib/Order/Interval/Set/Pi.lean @@ -99,7 +99,7 @@ theorem disjoint_pi_univ_Ioc_update_left_right {x y : ∀ i, α i} {i₀ : ι} { rw [disjoint_left] rintro z h₁ h₂ refine (h₁ i₀ (mem_univ _)).2.not_lt ?_ - simpa only [Function.update_same] using (h₂ i₀ (mem_univ _)).1 + simpa only [Function.update_self] using (h₂ i₀ (mem_univ _)).1 end PiPreorder @@ -115,11 +115,11 @@ theorem image_update_Icc (f : ∀ i, α i) (i : ι) (a b : α i) : refine ⟨?_, fun h => ⟨x i, ?_, ?_⟩⟩ · rintro ⟨c, hc, rfl⟩ simpa [update_le_update_iff] - · simpa only [Function.update_same] using h i (mem_univ i) + · simpa only [Function.update_self] using h i (mem_univ i) · ext j obtain rfl | hij := eq_or_ne i j - · exact Function.update_same _ _ _ - · simpa only [Function.update_noteq hij.symm, le_antisymm_iff] using h j (mem_univ j) + · exact Function.update_self .. + · simpa only [Function.update_of_ne hij.symm, le_antisymm_iff] using h j (mem_univ j) theorem image_update_Ico (f : ∀ i, α i) (i : ι) (a b : α i) : update f i '' Ico a b = Ico (update f i a) (update f i b) := by diff --git a/Mathlib/Order/Lattice.lean b/Mathlib/Order/Lattice.lean index acfa2bc47f35f..5cdba7b8a8b2a 100644 --- a/Mathlib/Order/Lattice.lean +++ b/Mathlib/Order/Lattice.lean @@ -856,11 +856,11 @@ variable {ι : Type*} {π : ι → Type*} [DecidableEq ι] -- Porting note: Dot notation on `Function.update` broke theorem update_sup [∀ i, SemilatticeSup (π i)] (f : ∀ i, π i) (i : ι) (a b : π i) : update f i (a ⊔ b) = update f i a ⊔ update f i b := - funext fun j => by obtain rfl | hji := eq_or_ne j i <;> simp [update_noteq, *] + funext fun j => by obtain rfl | hji := eq_or_ne j i <;> simp [update_of_ne, *] theorem update_inf [∀ i, SemilatticeInf (π i)] (f : ∀ i, π i) (i : ι) (a b : π i) : update f i (a ⊓ b) = update f i a ⊓ update f i b := - funext fun j => by obtain rfl | hji := eq_or_ne j i <;> simp [update_noteq, *] + funext fun j => by obtain rfl | hji := eq_or_ne j i <;> simp [update_of_ne, *] end Function diff --git a/Mathlib/Order/PiLex.lean b/Mathlib/Order/PiLex.lean index d58ef89f5bb1e..8dd412ba0637a 100644 --- a/Mathlib/Order/PiLex.lean +++ b/Mathlib/Order/PiLex.lean @@ -128,9 +128,9 @@ theorem lt_toLex_update_self_iff : toLex x < toLex (update x i a) ↔ x i < a := dsimp at h obtain rfl : j = i := by by_contra H - rw [update_noteq H] at h + rw [update_of_ne H] at h exact h.false - rwa [update_same] at h + rwa [update_self] at h @[simp] theorem toLex_update_lt_self_iff : toLex (update x i a) < toLex x ↔ a < x i := by @@ -139,9 +139,9 @@ theorem toLex_update_lt_self_iff : toLex (update x i a) < toLex x ↔ a < x i := dsimp at h obtain rfl : j = i := by by_contra H - rw [update_noteq H] at h + rw [update_of_ne H] at h exact h.false - rwa [update_same] at h + rwa [update_self] at h @[simp] theorem le_toLex_update_self_iff : toLex x ≤ toLex (update x i a) ↔ x i ≤ a := by @@ -176,10 +176,10 @@ instance [Preorder ι] [∀ i, LT (β i)] [∀ i, DenselyOrdered (β i)] : refine ⟨Function.update a₂ _ a, ⟨i, fun j hj => ?_, ?_⟩, i, fun j hj => ?_, ?_⟩ · rw [h j hj] dsimp only at hj - rw [Function.update_noteq hj.ne a] - · rwa [Function.update_same i a] - · rw [Function.update_noteq hj.ne a] - · rwa [Function.update_same i a]⟩ + rw [Function.update_of_ne hj.ne a] + · rwa [Function.update_self i a] + · rw [Function.update_of_ne hj.ne a] + · rwa [Function.update_self i a]⟩ theorem Lex.noMaxOrder' [Preorder ι] [∀ i, LT (β i)] (i : ι) [NoMaxOrder (β i)] : NoMaxOrder (Lex (∀ i, β i)) := @@ -187,7 +187,7 @@ theorem Lex.noMaxOrder' [Preorder ι] [∀ i, LT (β i)] (i : ι) [NoMaxOrder ( let ⟨b, hb⟩ := exists_gt (a i) classical exact ⟨Function.update a i b, i, fun j hj => - (Function.update_noteq hj.ne b a).symm, by rwa [Function.update_same i b]⟩⟩ + (Function.update_of_ne hj.ne b a).symm, by rwa [Function.update_self i b]⟩⟩ instance [LinearOrder ι] [WellFoundedLT ι] [Nonempty ι] [∀ i, PartialOrder (β i)] [∀ i, NoMaxOrder (β i)] : NoMaxOrder (Lex (∀ i, β i)) := diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index b6fe6957a17a2..95aa7008f3ea2 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -328,11 +328,11 @@ theorem IsMaximal.comap_piEvalRingHom {ι : Type*} {R : ι → Type*} [∀ i, Se have ⟨r, y, hy, eq⟩ := h.exists_inv hxI classical convert J.add_mem (J.mul_mem_left (update 0 i r) hxJ) - (b := update 1 i y) (le <| by apply update_same i y 1 ▸ hy) + (b := update 1 i y) (le <| by apply update_self i y 1 ▸ hy) ext j obtain rfl | ne := eq_or_ne j i · simpa [eq_comm] using eq - · simp [update_noteq ne] + · simp [update_of_ne ne] end Surjective diff --git a/Mathlib/RingTheory/Localization/Basic.lean b/Mathlib/RingTheory/Localization/Basic.lean index 6f9bf16ab95c9..22ac6c09fc61e 100644 --- a/Mathlib/RingTheory/Localization/Basic.lean +++ b/Mathlib/RingTheory/Localization/Basic.lean @@ -93,12 +93,12 @@ theorem mapPiEvalRingHom_bijective : Bijective (mapPiEvalRingHom S) := by simp_rw [map_mk'] at eq rw [IsLocalization.eq] at eq ⊢ obtain ⟨s, hs⟩ := eq - refine ⟨⟨update 0 i s, by apply update_same i s.1 0 ▸ s.2⟩, funext fun j ↦ ?_⟩ + refine ⟨⟨update 0 i s, by apply update_self i s.1 0 ▸ s.2⟩, funext fun j ↦ ?_⟩ obtain rfl | ne := eq_or_ne j i · simpa using hs - · simp [update_noteq ne] + · simp [update_of_ne ne] · obtain ⟨r, s, rfl⟩ := mk'_surjective S x - exact ⟨mk' (M := T) _ (update 0 i r) ⟨update 0 i s, by apply update_same i s.1 0 ▸ s.2⟩, + exact ⟨mk' (M := T) _ (update 0 i r) ⟨update 0 i s, by apply update_self i s.1 0 ▸ s.2⟩, by simp [map_mk']⟩ end Localization diff --git a/Mathlib/RingTheory/MaximalSpectrum.lean b/Mathlib/RingTheory/MaximalSpectrum.lean index 9e400e2b5fb1c..8db93538adf3f 100644 --- a/Mathlib/RingTheory/MaximalSpectrum.lean +++ b/Mathlib/RingTheory/MaximalSpectrum.lean @@ -160,10 +160,10 @@ theorem toPiLocalization_not_surjective_of_infinite [Infinite ι] : dsimp only [toPiLocalization_apply_apply, Subtype.coe_mk] at hr simp_rw [toPiLocalization_apply_apply, ← Localization.AtPrime.mapPiEvalRingHom_algebraMap_apply, hr] - rw [Function.update_noteq]; · simp_rw [Pi.zero_apply, map_zero] + rw [Function.update_of_ne]; · simp_rw [Pi.zero_apply, map_zero] exact fun h ↦ nmem ⟨⟨i, I.1, I.2.isPrime⟩, PrimeSpectrum.ext congr($h.1)⟩ replace hr := congr_fun hr ⟨J, max⟩ - rw [this, map_zero, Function.update_same] at hr + rw [this, map_zero, Function.update_self] at hr exact zero_ne_one hr variable {R} @@ -234,10 +234,10 @@ theorem isMaximal_of_toPiLocalization_surjective (surj : Function.Surjective (to have ⟨J, max, le⟩ := I.1.exists_le_maximal I.2.ne_top obtain ⟨r, hr⟩ := surj (Function.update 0 ⟨J, max.isPrime⟩ 1) by_contra h - have hJ : algebraMap _ _ r = _ := (congr_fun hr _).trans (Function.update_same _ _ _) + have hJ : algebraMap _ _ r = _ := (congr_fun hr _).trans (Function.update_self ..) have hI : algebraMap _ _ r = _ := congr_fun hr I rw [← IsLocalization.lift_eq (M := J.primeCompl) (S := Localization J.primeCompl), hJ, map_one, - Function.update_noteq] at hI + Function.update_of_ne] at hI · exact one_ne_zero hI · intro eq; have : I.1 = J := congr_arg (·.1) eq; exact h (this ▸ max) · exact fun ⟨s, hs⟩ ↦ IsLocalization.map_units (M := I.1.primeCompl) _ ⟨s, fun h ↦ hs (le h)⟩ diff --git a/Mathlib/RingTheory/MvPowerSeries/Basic.lean b/Mathlib/RingTheory/MvPowerSeries/Basic.lean index 481c58942c5c4..136414a51996a 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Basic.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Basic.lean @@ -641,12 +641,12 @@ theorem coeff_prod [DecidableEq σ] rintro u v rfl rw [ih, Finset.mul_sum, ← Finset.sum_attach] apply Finset.sum_congr rfl - simp only [mem_attach, Finset.prod_insert ha, Function.update_same, forall_true_left, + simp only [mem_attach, Finset.prod_insert ha, Function.update_self, forall_true_left, Subtype.forall] rintro x - rw [Finset.prod_congr rfl] intro i hi - rw [Function.update_noteq] + rw [Function.update_of_ne] exact ne_of_mem_of_not_mem hi ha · simp only [Set.PairwiseDisjoint, Set.Pairwise, mem_coe, mem_antidiagonal, ne_eq, disjoint_left, mem_map, mem_attach, Function.Embedding.coeFn_mk, true_and, Subtype.exists, @@ -654,7 +654,7 @@ theorem coeff_prod [DecidableEq σ] Prod.forall, Prod.mk.injEq] rintro u v rfl u' v' huv h k - l - hkl obtain rfl : u' = u := by - simpa only [Finsupp.coe_update, Function.update_same] using DFunLike.congr_fun hkl a + simpa only [Finsupp.coe_update, Function.update_self] using DFunLike.congr_fun hkl a simp only [add_right_inj] at huv exact h rfl huv.symm diff --git a/Mathlib/RingTheory/PrimeSpectrum.lean b/Mathlib/RingTheory/PrimeSpectrum.lean index 53888f21eb0b4..cee8d1be56b21 100644 --- a/Mathlib/RingTheory/PrimeSpectrum.lean +++ b/Mathlib/RingTheory/PrimeSpectrum.lean @@ -606,7 +606,7 @@ theorem sigmaToPi_injective : (sigmaToPi R).Injective := fun ⟨i, p⟩ ⟨j, q simpa using congr_arg (Function.update (0 : ∀ i, R i) i x ∈ ·.asIdeal) eq · refine (p.1.ne_top_iff_one.mp p.2.ne_top ?_).elim have : Function.update (1 : ∀ i, R i) j 0 ∈ (sigmaToPi R ⟨j, q⟩).asIdeal := by simp - simpa [← eq, Function.update_noteq ne] + simpa [← eq, Function.update_of_ne ne] variable [Infinite ι] [∀ i, Nontrivial (R i)] diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean index 937a415544928..2e6821abc2e9b 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean @@ -70,9 +70,9 @@ theorem HasProd.update (hf : HasProd f a₁) (b : β) [DecidableEq β] (a : α) HasProd (update f b a) (a / f b * a₁) := by convert (hasProd_ite_eq b (a / f b)).mul hf with b' by_cases h : b' = b - · rw [h, update_same] + · rw [h, update_self] simp [eq_self_iff_true, if_true, sub_add_cancel] - · simp only [h, update_noteq, if_false, Ne, one_mul, not_false_iff] + · simp only [h, update_of_ne, if_false, Ne, one_mul, not_false_iff] @[to_additive] theorem Multipliable.update (hf : Multipliable f) (b : β) [DecidableEq β] (a : α) : diff --git a/Mathlib/Topology/Algebra/Module/Multilinear/Bounded.lean b/Mathlib/Topology/Algebra/Module/Multilinear/Bounded.lean index cc362b7850cfb..3533ced7229f2 100644 --- a/Mathlib/Topology/Algebra/Module/Multilinear/Bounded.lean +++ b/Mathlib/Topology/Algebra/Module/Multilinear/Bounded.lean @@ -69,14 +69,14 @@ theorem image_multilinear' [Nonempty ι] {s : Set (∀ i, E i)} (hs : IsVonNBoun calc f (update y i₀ ((a / ∏ i ∈ I, c i) • y i₀)) ∈ V := hft fun i hi => by rcases eq_or_ne i i₀ with rfl | hne - · simp_rw [update_same, y, I.piecewise_eq_of_mem _ _ hi, smul_smul] + · simp_rw [update_self, y, I.piecewise_eq_of_mem _ _ hi, smul_smul] refine hc _ _ ?_ _ hx calc ‖(a / ∏ i ∈ I, c i) * c i‖ ≤ (‖∏ i ∈ I, c i‖ / ‖∏ i ∈ I, c i‖) * ‖c i‖ := by rw [norm_mul, norm_div]; gcongr; exact ha.out.le _ ≤ 1 * ‖c i‖ := by gcongr; apply div_self_le_one _ = ‖c i‖ := one_mul _ - · simp_rw [update_noteq hne, y, I.piecewise_eq_of_mem _ _ hi] + · simp_rw [update_of_ne hne, y, I.piecewise_eq_of_mem _ _ hi] exact hc _ _ le_rfl _ hx _ = a • f x := by rw [f.map_update_smul, update_eq_self, f.map_piecewise_smul, div_eq_mul_inv, mul_smul, diff --git a/Mathlib/Topology/Algebra/WithZeroTopology.lean b/Mathlib/Topology/Algebra/WithZeroTopology.lean index cd47b1c9d0f8f..461d02ba18734 100644 --- a/Mathlib/Topology/Algebra/WithZeroTopology.lean +++ b/Mathlib/Topology/Algebra/WithZeroTopology.lean @@ -51,7 +51,7 @@ theorem nhds_eq_update : (𝓝 : Γ₀ → Filter Γ₀) = update pure 0 (⨅ γ -/ theorem nhds_zero : 𝓝 (0 : Γ₀) = ⨅ γ ≠ 0, 𝓟 (Iio γ) := by - rw [nhds_eq_update, update_same] + rw [nhds_eq_update, update_self] /-- In a linearly ordered group with zero element adjoined, `U` is a neighbourhood of `0` if and only if there exists a nonzero element `γ₀` such that `Iio γ₀ ⊆ U`. -/ diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index 56c76196a7d53..a919902aa224f 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -1296,10 +1296,10 @@ theorem continuousWithinAt_update_same [DecidableEq α] {y : β} : ContinuousWithinAt (update f x y) s x ↔ Tendsto f (𝓝[s \ {x}] x) (𝓝 y) := calc ContinuousWithinAt (update f x y) s x ↔ Tendsto (update f x y) (𝓝[s \ {x}] x) (𝓝 y) := by - { rw [← continuousWithinAt_diff_self, ContinuousWithinAt, update_same] } + { rw [← continuousWithinAt_diff_self, ContinuousWithinAt, update_self] } _ ↔ Tendsto f (𝓝[s \ {x}] x) (𝓝 y) := tendsto_congr' <| eventually_nhdsWithin_iff.2 <| Eventually.of_forall - fun _ hz => update_noteq hz.2 _ _ + fun _ hz => update_of_ne hz.2 .. @[simp] theorem continuousAt_update_same [DecidableEq α] {y : β} : diff --git a/Mathlib/Topology/Order.lean b/Mathlib/Topology/Order.lean index 168716335f5dc..0be6eb00b68ce 100644 --- a/Mathlib/Topology/Order.lean +++ b/Mathlib/Topology/Order.lean @@ -123,8 +123,8 @@ theorem nhds_mkOfNhds_single [DecidableEq α] {a₀ : α} {l : Filter α} (h : p · filter_upwards [hs] with b hb rcases eq_or_ne b a with (rfl | hb) · exact hs - · rwa [update_noteq hb] - · simpa only [update_noteq ha, mem_pure, eventually_pure] using hs + · rwa [update_of_ne hb] + · simpa only [update_of_ne ha, mem_pure, eventually_pure] using hs theorem nhds_mkOfNhds_filterBasis (B : α → FilterBasis α) (a : α) (h₀ : ∀ x, ∀ n ∈ B x, x ∈ n) (h₁ : ∀ x, ∀ n ∈ B x, ∃ n₁ ∈ B x, ∀ x' ∈ n₁, ∃ n₂ ∈ B x', n₂ ⊆ n) : diff --git a/Mathlib/Topology/Separation/Basic.lean b/Mathlib/Topology/Separation/Basic.lean index 9f6cb623f903b..b49a1252d44b7 100644 --- a/Mathlib/Topology/Separation/Basic.lean +++ b/Mathlib/Topology/Separation/Basic.lean @@ -481,8 +481,8 @@ theorem continuousWithinAt_update_of_ne [T1Space X] [DecidableEq X] [Topological ContinuousWithinAt (Function.update f x y) s x' ↔ ContinuousWithinAt f s x' := EventuallyEq.congr_continuousWithinAt (mem_nhdsWithin_of_mem_nhds <| mem_of_superset (isOpen_ne.mem_nhds hne) fun _y' hy' => - Function.update_noteq hy' _ _) - (Function.update_noteq hne _ _) + Function.update_of_ne hy' _ _) + (Function.update_of_ne hne ..) theorem continuousAt_update_of_ne [T1Space X] [DecidableEq X] [TopologicalSpace Y] {f : X → Y} {x x' : X} {y : Y} (hne : x' ≠ x) : @@ -976,12 +976,12 @@ theorem IsCompact.finite_compact_cover {s : Set X} (hs : IsCompact s) {ι : Type refine ⟨update K x K₁, ?_, ?_, ?_⟩ · intro i rcases eq_or_ne i x with rfl | hi - · simp only [update_same, h1K₁] - · simp only [update_noteq hi, h1K] + · simp only [update_self, h1K₁] + · simp only [update_of_ne hi, h1K] · intro i rcases eq_or_ne i x with rfl | hi - · simp only [update_same, h2K₁] - · simp only [update_noteq hi, h2K] + · simp only [update_self, h2K₁] + · simp only [update_of_ne hi, h2K] · simp only [Finset.set_biUnion_insert_update _ hx, hK, h3K] theorem R1Space.of_continuous_specializes_imp [TopologicalSpace Y] {f : Y → X} (hc : Continuous f) diff --git a/Mathlib/Topology/ShrinkingLemma.lean b/Mathlib/Topology/ShrinkingLemma.lean index 2c85b7383164c..efcf49d420ddd 100644 --- a/Mathlib/Topology/ShrinkingLemma.lean +++ b/Mathlib/Topology/ShrinkingLemma.lean @@ -181,21 +181,21 @@ theorem exists_gt [NormalSpace X] (v : PartialRefinement u s ⊤) (hs : IsClosed · refine fun x hx => mem_iUnion.2 ?_ rcases em (∃ j ≠ i, x ∈ v j) with (⟨j, hji, hj⟩ | h) · use j - rwa [update_noteq hji] + rwa [update_of_ne hji] · push_neg at h use i - rw [update_same] + rw [update_self] exact hvi ⟨hx, mem_biInter h⟩ · rintro j (rfl | hj) - · rwa [update_same, ← v.apply_eq hi] - · rw [update_noteq (ne_of_mem_of_not_mem hj hi)] + · rwa [update_self, ← v.apply_eq hi] + · rw [update_of_ne (ne_of_mem_of_not_mem hj hi)] exact v.closure_subset hj · exact fun _ => trivial · intro j hj rw [mem_insert_iff, not_or] at hj - rw [update_noteq hj.1, v.apply_eq hj.2] + rw [update_of_ne hj.1, v.apply_eq hj.2] · refine ⟨subset_insert _ _, fun j hj => ?_⟩ - exact (update_noteq (ne_of_mem_of_not_mem hj hi) _ _).symm + exact (update_of_ne (ne_of_mem_of_not_mem hj hi) _ _).symm · exact fun hle => hi (hle.1 <| mem_insert _ _) end PartialRefinement @@ -302,37 +302,37 @@ theorem exists_gt_t2space (v : PartialRefinement u s (fun w => IsCompact (closur · refine fun x hx => mem_iUnion.2 ?_ rcases em (∃ j ≠ i, x ∈ v j) with (⟨j, hji, hj⟩ | h) · use j - rwa [update_noteq hji] + rwa [update_of_ne hji] · push_neg at h use i - rw [update_same] + rw [update_self] apply hvi.2.1 rw [hsi] exact ⟨hx, mem_iInter₂_of_mem h⟩ · rintro j (rfl | hj) - · rw [update_same] + · rw [update_self] exact subset_trans hvi.2.2.1 <| PartialRefinement.subset v j - · rw [update_noteq (ne_of_mem_of_not_mem hj hi)] + · rw [update_of_ne (ne_of_mem_of_not_mem hj hi)] exact v.closure_subset hj · intro j hj rw [mem_insert_iff] at hj by_cases h : j = i · rw [← h] - simp only [update_same] + simp only [update_self] exact hvi.2.2.2 · apply hj.elim · intro hji exact False.elim (h hji) · intro hjmemv - rw [update_noteq h] + rw [update_of_ne h] exact v.pred_of_mem hjmemv · intro j hj rw [mem_insert_iff, not_or] at hj - rw [update_noteq hj.1, v.apply_eq hj.2] + rw [update_of_ne hj.1, v.apply_eq hj.2] · refine ⟨subset_insert _ _, fun j hj => ?_⟩ - exact (update_noteq (ne_of_mem_of_not_mem hj hi) _ _).symm + exact (update_of_ne (ne_of_mem_of_not_mem hj hi) _ _).symm · exact fun hle => hi (hle.1 <| mem_insert _ _) - · simp only [update_same] + · simp only [update_self] exact hvi.2.2.2 /-- **Shrinking lemma** . A point-finite open cover of a compact subset of a `T2Space` From e423a1f0105a779265101af2bd177f5d8065c594 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 29 Dec 2024 04:54:57 +0000 Subject: [PATCH 026/189] chore(Calculus): group together "converse to the mean value inequality" theorems (#20307) Generalizing this file is challenging because the lemmas are in an arbitrary order. This brings the order of the FDeriv file in line with the Deriv file, and adds matching section names to try and prevent these separating in future. --- Mathlib/Analysis/Calculus/Deriv/Basic.lean | 4 + Mathlib/Analysis/Calculus/FDeriv/Basic.lean | 122 ++++++++++---------- 2 files changed, 66 insertions(+), 60 deletions(-) diff --git a/Mathlib/Analysis/Calculus/Deriv/Basic.lean b/Mathlib/Analysis/Calculus/Deriv/Basic.lean index 48874c1c046ed..d0e6db0c4bcf9 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Basic.lean @@ -680,6 +680,8 @@ protected theorem HasDerivAt.continuousOn {f f' : 𝕜 → F} (hderiv : ∀ x end Continuous +section MeanValue + /-- Converse to the mean value inequality: if `f` is differentiable at `x₀` and `C`-lipschitz on a neighborhood of `x₀` then its derivative at `x₀` has norm bounded by `C`. This version only assumes that `‖f x - f x₀‖ ≤ C * ‖x - x₀‖` in a neighborhood of `x`. -/ @@ -721,3 +723,5 @@ Version using `deriv`. -/ theorem norm_deriv_le_of_lipschitz {f : 𝕜 → F} {x₀ : 𝕜} {C : ℝ≥0} (hlip : LipschitzWith C f) : ‖deriv f x₀‖ ≤ C := by simpa [norm_deriv_eq_norm_fderiv] using norm_fderiv_le_of_lipschitz 𝕜 hlip + +end MeanValue diff --git a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean index 101108c7507b4..ee68296873c8a 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean @@ -335,35 +335,6 @@ theorem hasFDerivAt_iff_isLittleO_nhds_zero : rw [HasFDerivAt, hasFDerivAtFilter_iff_isLittleO, ← map_add_left_nhds_zero x, isLittleO_map] simp [Function.comp_def] -/-- Converse to the mean value inequality: if `f` is differentiable at `x₀` and `C`-lipschitz -on a neighborhood of `x₀` then its derivative at `x₀` has norm bounded by `C`. This version -only assumes that `‖f x - f x₀‖ ≤ C * ‖x - x₀‖` in a neighborhood of `x`. -/ -theorem HasFDerivAt.le_of_lip' {f : E → F} {f' : E →L[𝕜] F} {x₀ : E} (hf : HasFDerivAt f f' x₀) - {C : ℝ} (hC₀ : 0 ≤ C) (hlip : ∀ᶠ x in 𝓝 x₀, ‖f x - f x₀‖ ≤ C * ‖x - x₀‖) : ‖f'‖ ≤ C := by - refine le_of_forall_pos_le_add fun ε ε0 => opNorm_le_of_nhds_zero ?_ ?_ - · exact add_nonneg hC₀ ε0.le - rw [← map_add_left_nhds_zero x₀, eventually_map] at hlip - filter_upwards [isLittleO_iff.1 (hasFDerivAt_iff_isLittleO_nhds_zero.1 hf) ε0, hlip] with y hy hyC - rw [add_sub_cancel_left] at hyC - calc - ‖f' y‖ ≤ ‖f (x₀ + y) - f x₀‖ + ‖f (x₀ + y) - f x₀ - f' y‖ := norm_le_insert _ _ - _ ≤ C * ‖y‖ + ε * ‖y‖ := add_le_add hyC hy - _ = (C + ε) * ‖y‖ := (add_mul _ _ _).symm - -/-- Converse to the mean value inequality: if `f` is differentiable at `x₀` and `C`-lipschitz -on a neighborhood of `x₀` then its derivative at `x₀` has norm bounded by `C`. -/ -theorem HasFDerivAt.le_of_lipschitzOn - {f : E → F} {f' : E →L[𝕜] F} {x₀ : E} (hf : HasFDerivAt f f' x₀) - {s : Set E} (hs : s ∈ 𝓝 x₀) {C : ℝ≥0} (hlip : LipschitzOnWith C f s) : ‖f'‖ ≤ C := by - refine hf.le_of_lip' C.coe_nonneg ?_ - filter_upwards [hs] with x hx using hlip.norm_sub_le hx (mem_of_mem_nhds hs) - -/-- Converse to the mean value inequality: if `f` is differentiable at `x₀` and `C`-lipschitz -then its derivative at `x₀` has norm bounded by `C`. -/ -theorem HasFDerivAt.le_of_lipschitz {f : E → F} {f' : E →L[𝕜] F} {x₀ : E} (hf : HasFDerivAt f f' x₀) - {C : ℝ≥0} (hlip : LipschitzWith C f) : ‖f'‖ ≤ C := - hf.le_of_lipschitzOn univ_mem (lipschitzOnWith_univ.2 hlip) - nonrec theorem HasFDerivAtFilter.mono (h : HasFDerivAtFilter f f' x L₂) (hst : L₁ ≤ L₂) : HasFDerivAtFilter f f' x L₁ := .of_isLittleO <| h.isLittleO.mono hst @@ -559,37 +530,6 @@ protected theorem HasFDerivAt.fderiv (h : HasFDerivAt f f' x) : fderiv 𝕜 f x theorem fderiv_eq {f' : E → E →L[𝕜] F} (h : ∀ x, HasFDerivAt f (f' x) x) : fderiv 𝕜 f = f' := funext fun x => (h x).fderiv -variable (𝕜) - -/-- Converse to the mean value inequality: if `f` is `C`-lipschitz -on a neighborhood of `x₀` then its derivative at `x₀` has norm bounded by `C`. This version -only assumes that `‖f x - f x₀‖ ≤ C * ‖x - x₀‖` in a neighborhood of `x`. -/ -theorem norm_fderiv_le_of_lip' {f : E → F} {x₀ : E} - {C : ℝ} (hC₀ : 0 ≤ C) (hlip : ∀ᶠ x in 𝓝 x₀, ‖f x - f x₀‖ ≤ C * ‖x - x₀‖) : - ‖fderiv 𝕜 f x₀‖ ≤ C := by - by_cases hf : DifferentiableAt 𝕜 f x₀ - · exact hf.hasFDerivAt.le_of_lip' hC₀ hlip - · rw [fderiv_zero_of_not_differentiableAt hf] - simp [hC₀] - -/-- Converse to the mean value inequality: if `f` is `C`-lipschitz -on a neighborhood of `x₀` then its derivative at `x₀` has norm bounded by `C`. -Version using `fderiv`. -/ --- Porting note: renamed so that dot-notation makes sense -theorem norm_fderiv_le_of_lipschitzOn {f : E → F} {x₀ : E} {s : Set E} (hs : s ∈ 𝓝 x₀) - {C : ℝ≥0} (hlip : LipschitzOnWith C f s) : ‖fderiv 𝕜 f x₀‖ ≤ C := by - refine norm_fderiv_le_of_lip' 𝕜 C.coe_nonneg ?_ - filter_upwards [hs] with x hx using hlip.norm_sub_le hx (mem_of_mem_nhds hs) - -/-- Converse to the mean value inequality: if `f` is `C`-lipschitz then -its derivative at `x₀` has norm bounded by `C`. -Version using `fderiv`. -/ -theorem norm_fderiv_le_of_lipschitz {f : E → F} {x₀ : E} - {C : ℝ≥0} (hlip : LipschitzWith C f) : ‖fderiv 𝕜 f x₀‖ ≤ C := - norm_fderiv_le_of_lipschitzOn 𝕜 univ_mem (lipschitzOnWith_univ.2 hlip) - -variable {𝕜} - protected theorem HasFDerivWithinAt.fderivWithin (h : HasFDerivWithinAt f f' s x) (hxs : UniqueDiffWithinAt 𝕜 s x) : fderivWithin 𝕜 f s x = f' := (hxs.eq h h.differentiableWithinAt.hasFDerivWithinAt).symm @@ -1148,6 +1088,68 @@ theorem hasFDerivAt_zero_of_eventually_const (c : F) (hf : f =ᶠ[𝓝 x] fun _ end Const +section MeanValue + +/-- Converse to the mean value inequality: if `f` is differentiable at `x₀` and `C`-lipschitz +on a neighborhood of `x₀` then its derivative at `x₀` has norm bounded by `C`. This version +only assumes that `‖f x - f x₀‖ ≤ C * ‖x - x₀‖` in a neighborhood of `x`. -/ +theorem HasFDerivAt.le_of_lip' {f : E → F} {f' : E →L[𝕜] F} {x₀ : E} (hf : HasFDerivAt f f' x₀) + {C : ℝ} (hC₀ : 0 ≤ C) (hlip : ∀ᶠ x in 𝓝 x₀, ‖f x - f x₀‖ ≤ C * ‖x - x₀‖) : ‖f'‖ ≤ C := by + refine le_of_forall_pos_le_add fun ε ε0 => opNorm_le_of_nhds_zero ?_ ?_ + · exact add_nonneg hC₀ ε0.le + rw [← map_add_left_nhds_zero x₀, eventually_map] at hlip + filter_upwards [isLittleO_iff.1 (hasFDerivAt_iff_isLittleO_nhds_zero.1 hf) ε0, hlip] with y hy hyC + rw [add_sub_cancel_left] at hyC + calc + ‖f' y‖ ≤ ‖f (x₀ + y) - f x₀‖ + ‖f (x₀ + y) - f x₀ - f' y‖ := norm_le_insert _ _ + _ ≤ C * ‖y‖ + ε * ‖y‖ := add_le_add hyC hy + _ = (C + ε) * ‖y‖ := (add_mul _ _ _).symm + +/-- Converse to the mean value inequality: if `f` is differentiable at `x₀` and `C`-lipschitz +on a neighborhood of `x₀` then its derivative at `x₀` has norm bounded by `C`. -/ +theorem HasFDerivAt.le_of_lipschitzOn + {f : E → F} {f' : E →L[𝕜] F} {x₀ : E} (hf : HasFDerivAt f f' x₀) + {s : Set E} (hs : s ∈ 𝓝 x₀) {C : ℝ≥0} (hlip : LipschitzOnWith C f s) : ‖f'‖ ≤ C := by + refine hf.le_of_lip' C.coe_nonneg ?_ + filter_upwards [hs] with x hx using hlip.norm_sub_le hx (mem_of_mem_nhds hs) + +/-- Converse to the mean value inequality: if `f` is differentiable at `x₀` and `C`-lipschitz +then its derivative at `x₀` has norm bounded by `C`. -/ +theorem HasFDerivAt.le_of_lipschitz {f : E → F} {f' : E →L[𝕜] F} {x₀ : E} (hf : HasFDerivAt f f' x₀) + {C : ℝ≥0} (hlip : LipschitzWith C f) : ‖f'‖ ≤ C := + hf.le_of_lipschitzOn univ_mem (lipschitzOnWith_univ.2 hlip) + +variable (𝕜) + +/-- Converse to the mean value inequality: if `f` is `C`-lipschitz +on a neighborhood of `x₀` then its derivative at `x₀` has norm bounded by `C`. This version +only assumes that `‖f x - f x₀‖ ≤ C * ‖x - x₀‖` in a neighborhood of `x`. -/ +theorem norm_fderiv_le_of_lip' {f : E → F} {x₀ : E} + {C : ℝ} (hC₀ : 0 ≤ C) (hlip : ∀ᶠ x in 𝓝 x₀, ‖f x - f x₀‖ ≤ C * ‖x - x₀‖) : + ‖fderiv 𝕜 f x₀‖ ≤ C := by + by_cases hf : DifferentiableAt 𝕜 f x₀ + · exact hf.hasFDerivAt.le_of_lip' hC₀ hlip + · rw [fderiv_zero_of_not_differentiableAt hf] + simp [hC₀] + +/-- Converse to the mean value inequality: if `f` is `C`-lipschitz +on a neighborhood of `x₀` then its derivative at `x₀` has norm bounded by `C`. +Version using `fderiv`. -/ +-- Porting note: renamed so that dot-notation makes sense +theorem norm_fderiv_le_of_lipschitzOn {f : E → F} {x₀ : E} {s : Set E} (hs : s ∈ 𝓝 x₀) + {C : ℝ≥0} (hlip : LipschitzOnWith C f s) : ‖fderiv 𝕜 f x₀‖ ≤ C := by + refine norm_fderiv_le_of_lip' 𝕜 C.coe_nonneg ?_ + filter_upwards [hs] with x hx using hlip.norm_sub_le hx (mem_of_mem_nhds hs) + +/-- Converse to the mean value inequality: if `f` is `C`-lipschitz then +its derivative at `x₀` has norm bounded by `C`. +Version using `fderiv`. -/ +theorem norm_fderiv_le_of_lipschitz {f : E → F} {x₀ : E} + {C : ℝ≥0} (hlip : LipschitzWith C f) : ‖fderiv 𝕜 f x₀‖ ≤ C := + norm_fderiv_le_of_lipschitzOn 𝕜 univ_mem (lipschitzOnWith_univ.2 hlip) + +end MeanValue + end /-! ### Support of derivatives -/ From 0da6485fb4629890947f49ed3ee5e3d5269ff32c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 29 Dec 2024 09:21:07 +0000 Subject: [PATCH 027/189] =?UTF-8?q?feat:=20`fun=20n=20:=20=E2=84=95=20?= =?UTF-8?q?=E2=86=A6=20n=20^=20n`=20is=20monotone=20(#20296)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From GrowthInGroups (LeanCamCombi) --- Mathlib/Order/Monotone/Basic.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/Order/Monotone/Basic.lean b/Mathlib/Order/Monotone/Basic.lean index 163c8a7c18ccd..4210837c3e715 100644 --- a/Mathlib/Order/Monotone/Basic.lean +++ b/Mathlib/Order/Monotone/Basic.lean @@ -952,6 +952,18 @@ function `ℕ → α`. -/ theorem exists_strictAnti [Nonempty α] [NoMinOrder α] : ∃ f : ℕ → α, StrictAnti f := exists_strictMono αᵒᵈ +lemma pow_self_mono : Monotone fun n : ℕ ↦ n ^ n := by + refine monotone_nat_of_le_succ fun n ↦ ?_ + rw [Nat.pow_succ] + exact (Nat.pow_le_pow_left n.le_succ _).trans (Nat.le_mul_of_pos_right _ n.succ_pos) + +lemma pow_monotoneOn : MonotoneOn (fun p : ℕ × ℕ ↦ p.1 ^ p.2) {p | p.1 ≠ 0} := fun _p _ _q hq hpq ↦ + (Nat.pow_le_pow_left hpq.1 _).trans (Nat.pow_le_pow_right (Nat.pos_iff_ne_zero.2 hq) hpq.2) + +lemma pow_self_strictMonoOn : StrictMonoOn (fun n : ℕ ↦ n ^ n) {n : ℕ | n ≠ 0} := + fun _m hm _n hn hmn ↦ + (Nat.pow_lt_pow_left hmn hm).trans_le (Nat.pow_le_pow_right (Nat.pos_iff_ne_zero.2 hn) hmn.le) + end Nat theorem Int.rel_of_forall_rel_succ_of_lt (r : β → β → Prop) [IsTrans β r] {f : ℤ → β} From e785d99ad9a9af18b3ed0e1ed837a0b253df7a38 Mon Sep 17 00:00:00 2001 From: smorel394 Date: Sun, 29 Dec 2024 10:18:52 +0000 Subject: [PATCH 028/189] chore(CategoryTheory/Pretriangulated/Opposite): split CategoryTheory.Triangulated.Opposite into three files (#20316) This splits the long (and getting longer with a new PR) file `CatgeoryTheory.Triangulated.Opposite` in three files: - `Triangulated.Opposite.Basic`: just the shift on the opposite category + lemmas - `Triangulated.Opposite.Triangle`: the equivalence `triangleOpEquivalence` - `Triangulated.Opposite.Pretriangulated`: the pretriangulated structure --- Mathlib.lean | 4 +- .../DerivedCategory/Ext/ExactSequences.lean | 2 - .../Shift/ShiftedHomOpposite.lean | 2 +- .../{Opposite.lean => Opposite/Basic.lean} | 252 +----------------- .../Opposite/Pretriangulated.lean | 195 ++++++++++++++ .../Triangulated/Opposite/Triangle.lean | 110 ++++++++ .../CategoryTheory/Triangulated/Yoneda.lean | 2 +- 7 files changed, 316 insertions(+), 251 deletions(-) rename Mathlib/CategoryTheory/Triangulated/{Opposite.lean => Opposite/Basic.lean} (50%) create mode 100644 Mathlib/CategoryTheory/Triangulated/Opposite/Pretriangulated.lean create mode 100644 Mathlib/CategoryTheory/Triangulated/Opposite/Triangle.lean diff --git a/Mathlib.lean b/Mathlib.lean index d6a597218d32d..1625c95f15a04 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2205,7 +2205,9 @@ import Mathlib.CategoryTheory.Thin import Mathlib.CategoryTheory.Triangulated.Basic import Mathlib.CategoryTheory.Triangulated.Functor import Mathlib.CategoryTheory.Triangulated.HomologicalFunctor -import Mathlib.CategoryTheory.Triangulated.Opposite +import Mathlib.CategoryTheory.Triangulated.Opposite.Basic +import Mathlib.CategoryTheory.Triangulated.Opposite.Pretriangulated +import Mathlib.CategoryTheory.Triangulated.Opposite.Triangle import Mathlib.CategoryTheory.Triangulated.Pretriangulated import Mathlib.CategoryTheory.Triangulated.Rotate import Mathlib.CategoryTheory.Triangulated.Subcategory diff --git a/Mathlib/Algebra/Homology/DerivedCategory/Ext/ExactSequences.lean b/Mathlib/Algebra/Homology/DerivedCategory/Ext/ExactSequences.lean index cee76c6862e5c..a8062d76e37e9 100644 --- a/Mathlib/Algebra/Homology/DerivedCategory/Ext/ExactSequences.lean +++ b/Mathlib/Algebra/Homology/DerivedCategory/Ext/ExactSequences.lean @@ -3,9 +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.Algebra.Exact import Mathlib.Algebra.Homology.DerivedCategory.Ext.ExtClass -import Mathlib.Algebra.Homology.ShortComplex.Ab import Mathlib.CategoryTheory.Triangulated.Yoneda /-! diff --git a/Mathlib/CategoryTheory/Shift/ShiftedHomOpposite.lean b/Mathlib/CategoryTheory/Shift/ShiftedHomOpposite.lean index 5f21658547de3..95b881bbdf2f9 100644 --- a/Mathlib/CategoryTheory/Shift/ShiftedHomOpposite.lean +++ b/Mathlib/CategoryTheory/Shift/ShiftedHomOpposite.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.CategoryTheory.Triangulated.Opposite +import Mathlib.CategoryTheory.Triangulated.Opposite.Basic import Mathlib.CategoryTheory.Shift.ShiftedHom /-! Shifted morphisms in the opposite category diff --git a/Mathlib/CategoryTheory/Triangulated/Opposite.lean b/Mathlib/CategoryTheory/Triangulated/Opposite/Basic.lean similarity index 50% rename from Mathlib/CategoryTheory/Triangulated/Opposite.lean rename to Mathlib/CategoryTheory/Triangulated/Opposite/Basic.lean index 9e34b1cc2e37f..94ae9013b1c47 100644 --- a/Mathlib/CategoryTheory/Triangulated/Opposite.lean +++ b/Mathlib/CategoryTheory/Triangulated/Opposite/Basic.lean @@ -3,18 +3,17 @@ Copyright (c) 2023 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.RegularMono import Mathlib.CategoryTheory.Shift.Opposite import Mathlib.CategoryTheory.Shift.Pullback -import Mathlib.CategoryTheory.Triangulated.HomologicalFunctor -import Mathlib.Tactic.Linarith /-! -# The (pre)triangulated structure on the opposite category +# The shift on the opposite category of a pretriangulated category -In this file, we shall construct the (pre)triangulated structure -on the opposite category `Cᵒᵖ` of a (pre)triangulated category `C`. - -The shift on `Cᵒᵖ` is obtained by combining the constructions in the files +Let `C` be a (pre)triangulated category. We already have a shift on `Cᵒᵖ` given +by `CategoryTheory.Shift.Opposite`, but this is not the shift that we want to +make `Cᵒᵖ` into a (pre)triangulated category. +The correct shift on `Cᵒᵖ` is obtained by combining the constructions in the files `CategoryTheory.Shift.Opposite` and `CategoryTheory.Shift.Pullback`. When the user opens `CategoryTheory.Pretriangulated.Opposite`, the category `Cᵒᵖ` is equipped with the shift by `ℤ` such that @@ -28,12 +27,6 @@ Some compatibilities between the shifts on `C` and `Cᵒᵖ` are also expressed the equivalence of categories `opShiftFunctorEquivalence C n : Cᵒᵖ ≌ Cᵒᵖ` whose functor is `shiftFunctor Cᵒᵖ n` and whose inverse functor is `(shiftFunctor C n).op`. -If `X ⟶ Y ⟶ Z ⟶ X⟦1⟧` is a distinguished triangle in `C`, then the triangle -`op Z ⟶ op Y ⟶ op X ⟶ (op Z)⟦1⟧` that is deduced *without introducing signs* -shall be a distinguished triangle in `Cᵒᵖ`. This is equivalent to the definition -in [Verdiers's thesis, p. 96][verdier1996] which would require that the triangle -`(op X)⟦-1⟧ ⟶ op Z ⟶ op Y ⟶ op X` (without signs) is *antidistinguished*. - ## References * [Jean-Louis Verdier, *Des catégories dérivées des catégories abéliennes*][verdier1996] @@ -248,239 +241,6 @@ lemma opShiftFunctorEquivalence_counitIso_hom_app_shift (X : Cᵒᵖ) (n : ℤ) ((opShiftFunctorEquivalence C n).unitIso.inv.app X)⟦n⟧' := (opShiftFunctorEquivalence C n).counit_app_functor X -variable (C) - -namespace TriangleOpEquivalence - -/-- The functor which sends a triangle `X ⟶ Y ⟶ Z ⟶ X⟦1⟧` in `C` to the triangle -`op Z ⟶ op Y ⟶ op X ⟶ (op Z)⟦1⟧` in `Cᵒᵖ` (without introducing signs). -/ -@[simps] -noncomputable def functor : (Triangle C)ᵒᵖ ⥤ Triangle Cᵒᵖ where - obj T := Triangle.mk T.unop.mor₂.op T.unop.mor₁.op - ((opShiftFunctorEquivalence C 1).counitIso.inv.app (Opposite.op T.unop.obj₁) ≫ - T.unop.mor₃.op⟦(1 : ℤ)⟧') - map {T₁ T₂} φ := - { hom₁ := φ.unop.hom₃.op - hom₂ := φ.unop.hom₂.op - hom₃ := φ.unop.hom₁.op - comm₁ := Quiver.Hom.unop_inj φ.unop.comm₂.symm - comm₂ := Quiver.Hom.unop_inj φ.unop.comm₁.symm - comm₃ := by - dsimp - rw [assoc, ← Functor.map_comp, ← op_comp, ← φ.unop.comm₃, op_comp, Functor.map_comp, - opShiftFunctorEquivalence_counitIso_inv_naturality_assoc] - rfl } - -/-- The functor which sends a triangle `X ⟶ Y ⟶ Z ⟶ X⟦1⟧` in `Cᵒᵖ` to the triangle -`Z.unop ⟶ Y.unop ⟶ X.unop ⟶ Z.unop⟦1⟧` in `C` (without introducing signs). -/ -@[simps] -noncomputable def inverse : Triangle Cᵒᵖ ⥤ (Triangle C)ᵒᵖ where - obj T := Opposite.op (Triangle.mk T.mor₂.unop T.mor₁.unop - (((opShiftFunctorEquivalence C 1).unitIso.inv.app T.obj₁).unop ≫ T.mor₃.unop⟦(1 : ℤ)⟧')) - map {T₁ T₂} φ := Quiver.Hom.op - { hom₁ := φ.hom₃.unop - hom₂ := φ.hom₂.unop - hom₃ := φ.hom₁.unop - comm₁ := Quiver.Hom.op_inj φ.comm₂.symm - comm₂ := Quiver.Hom.op_inj φ.comm₁.symm - comm₃ := Quiver.Hom.op_inj (by - dsimp - rw [assoc, ← opShiftFunctorEquivalence_unitIso_inv_naturality, - ← op_comp_assoc, ← Functor.map_comp, ← unop_comp, ← φ.comm₃, - unop_comp, Functor.map_comp, op_comp, assoc]) } - -/-- The unit isomorphism of the -equivalence `triangleOpEquivalence C : (Triangle C)ᵒᵖ ≌ Triangle Cᵒᵖ` . -/ -@[simps!] -noncomputable def unitIso : 𝟭 _ ≅ functor C ⋙ inverse C := - NatIso.ofComponents (fun T => Iso.op - (Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) (by aesop_cat) (by aesop_cat) - (Quiver.Hom.op_inj - (by simp [shift_unop_opShiftFunctorEquivalence_counitIso_inv_app])))) - (fun {T₁ T₂} f => Quiver.Hom.unop_inj (by aesop_cat)) - -/-- The counit isomorphism of the -equivalence `triangleOpEquivalence C : (Triangle C)ᵒᵖ ≌ Triangle Cᵒᵖ` . -/ -@[simps!] -noncomputable def counitIso : inverse C ⋙ functor C ≅ 𝟭 _ := - NatIso.ofComponents (fun T => by - refine Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) ?_ ?_ ?_ - · aesop_cat - · aesop_cat - · dsimp - rw [Functor.map_id, comp_id, id_comp, Functor.map_comp, - ← opShiftFunctorEquivalence_counitIso_inv_naturality_assoc, - opShiftFunctorEquivalence_counitIso_inv_app_shift, ← Functor.map_comp, - Iso.hom_inv_id_app, Functor.map_id] - simp only [Functor.id_obj, comp_id]) - (by aesop_cat) - -end TriangleOpEquivalence - -/-- An anti-equivalence between the categories of triangles in `C` and in `Cᵒᵖ`. -A triangle in `Cᵒᵖ` shall be distinguished iff it correspond to a distinguished -triangle in `C` via this equivalence. -/ -@[simps] -noncomputable def triangleOpEquivalence : - (Triangle C)ᵒᵖ ≌ Triangle Cᵒᵖ where - functor := TriangleOpEquivalence.functor C - inverse := TriangleOpEquivalence.inverse C - unitIso := TriangleOpEquivalence.unitIso C - counitIso := TriangleOpEquivalence.counitIso C - -variable [HasZeroObject C] [Preadditive C] [∀ (n : ℤ), (shiftFunctor C n).Additive] - [Pretriangulated C] - -namespace Opposite - -/-- A triangle in `Cᵒᵖ` shall be distinguished iff it corresponds to a distinguished -triangle in `C` via the equivalence `triangleOpEquivalence C : (Triangle C)ᵒᵖ ≌ Triangle Cᵒᵖ`. -/ -def distinguishedTriangles : Set (Triangle Cᵒᵖ) := - fun T => ((triangleOpEquivalence C).inverse.obj T).unop ∈ distTriang C - -variable {C} - -lemma mem_distinguishedTriangles_iff (T : Triangle Cᵒᵖ) : - T ∈ distinguishedTriangles C ↔ - ((triangleOpEquivalence C).inverse.obj T).unop ∈ distTriang C := by - rfl - -lemma mem_distinguishedTriangles_iff' (T : Triangle Cᵒᵖ) : - T ∈ distinguishedTriangles C ↔ - ∃ (T' : Triangle C) (_ : T' ∈ distTriang C), - Nonempty (T ≅ (triangleOpEquivalence C).functor.obj (Opposite.op T')) := by - rw [mem_distinguishedTriangles_iff] - constructor - · intro hT - exact ⟨_ ,hT, ⟨(triangleOpEquivalence C).counitIso.symm.app T⟩⟩ - · rintro ⟨T', hT', ⟨e⟩⟩ - refine isomorphic_distinguished _ hT' _ ?_ - exact Iso.unop ((triangleOpEquivalence C).unitIso.app (Opposite.op T') ≪≫ - (triangleOpEquivalence C).inverse.mapIso e.symm) - -lemma isomorphic_distinguished (T₁ : Triangle Cᵒᵖ) - (hT₁ : T₁ ∈ distinguishedTriangles C) (T₂ : Triangle Cᵒᵖ) (e : T₂ ≅ T₁) : - T₂ ∈ distinguishedTriangles C := by - simp only [mem_distinguishedTriangles_iff] at hT₁ ⊢ - exact Pretriangulated.isomorphic_distinguished _ hT₁ _ - ((triangleOpEquivalence C).inverse.mapIso e).unop.symm - -/-- Up to rotation, the contractible triangle `X ⟶ X ⟶ 0 ⟶ X⟦1⟧` for `X : Cᵒᵖ` corresponds -to the contractible triangle for `X.unop` in `C`. -/ -@[simps!] -noncomputable def contractibleTriangleIso (X : Cᵒᵖ) : - contractibleTriangle X ≅ (triangleOpEquivalence C).functor.obj - (Opposite.op (contractibleTriangle X.unop).invRotate) := - Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) - (IsZero.iso (isZero_zero _) (by - dsimp - rw [IsZero.iff_id_eq_zero] - change (𝟙 ((0 : C)⟦(-1 : ℤ)⟧)).op = 0 - rw [← Functor.map_id, id_zero, Functor.map_zero, op_zero])) - (by aesop_cat) (by aesop_cat) (by aesop_cat) - -lemma contractible_distinguished (X : Cᵒᵖ) : - contractibleTriangle X ∈ distinguishedTriangles C := by - rw [mem_distinguishedTriangles_iff'] - exact ⟨_, inv_rot_of_distTriang _ (Pretriangulated.contractible_distinguished X.unop), - ⟨contractibleTriangleIso X⟩⟩ - -/-- Isomorphism expressing a compatibility of the equivalence `triangleOpEquivalence C` -with the rotation of triangles. -/ -noncomputable def rotateTriangleOpEquivalenceInverseObjRotateUnopIso (T : Triangle Cᵒᵖ) : - ((triangleOpEquivalence C).inverse.obj T.rotate).unop.rotate ≅ - ((triangleOpEquivalence C).inverse.obj T).unop := - Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) - (-((opShiftFunctorEquivalence C 1).unitIso.app T.obj₁).unop) (by simp) - (Quiver.Hom.op_inj (by aesop_cat)) (by aesop_cat) - -lemma rotate_distinguished_triangle (T : Triangle Cᵒᵖ) : - T ∈ distinguishedTriangles C ↔ T.rotate ∈ distinguishedTriangles C := by - simp only [mem_distinguishedTriangles_iff, Pretriangulated.rotate_distinguished_triangle - ((triangleOpEquivalence C).inverse.obj (T.rotate)).unop] - exact distinguished_iff_of_iso (rotateTriangleOpEquivalenceInverseObjRotateUnopIso T).symm - -lemma distinguished_cocone_triangle {X Y : Cᵒᵖ} (f : X ⟶ Y) : - ∃ (Z : Cᵒᵖ) (g : Y ⟶ Z) (h : Z ⟶ X⟦(1 : ℤ)⟧), - Triangle.mk f g h ∈ distinguishedTriangles C := by - obtain ⟨Z, g, h, H⟩ := Pretriangulated.distinguished_cocone_triangle₁ f.unop - refine ⟨_, g.op, (opShiftFunctorEquivalence C 1).counitIso.inv.app (Opposite.op Z) ≫ - (shiftFunctor Cᵒᵖ (1 : ℤ)).map h.op, ?_⟩ - simp only [mem_distinguishedTriangles_iff] - refine Pretriangulated.isomorphic_distinguished _ H _ ?_ - exact Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) (by aesop_cat) (by aesop_cat) - (Quiver.Hom.op_inj (by simp [shift_unop_opShiftFunctorEquivalence_counitIso_inv_app])) - -lemma complete_distinguished_triangle_morphism (T₁ T₂ : Triangle Cᵒᵖ) - (hT₁ : T₁ ∈ distinguishedTriangles C) (hT₂ : T₂ ∈ distinguishedTriangles C) - (a : T₁.obj₁ ⟶ T₂.obj₁) (b : T₁.obj₂ ⟶ T₂.obj₂) (comm : T₁.mor₁ ≫ b = a ≫ T₂.mor₁) : - ∃ (c : T₁.obj₃ ⟶ T₂.obj₃), T₁.mor₂ ≫ c = b ≫ T₂.mor₂ ∧ - T₁.mor₃ ≫ a⟦1⟧' = c ≫ T₂.mor₃ := by - rw [mem_distinguishedTriangles_iff] at hT₁ hT₂ - obtain ⟨c, hc₁, hc₂⟩ := - Pretriangulated.complete_distinguished_triangle_morphism₁ _ _ hT₂ hT₁ - b.unop a.unop (Quiver.Hom.op_inj comm.symm) - dsimp at c hc₁ hc₂ - replace hc₂ := ((opShiftFunctorEquivalence C 1).unitIso.hom.app T₂.obj₁).unop ≫= hc₂ - dsimp at hc₂ - simp only [assoc, Iso.unop_hom_inv_id_app_assoc] at hc₂ - refine ⟨c.op, Quiver.Hom.unop_inj hc₁.symm, Quiver.Hom.unop_inj ?_⟩ - apply (shiftFunctor C (1 : ℤ)).map_injective - rw [unop_comp, unop_comp, Functor.map_comp, Functor.map_comp, - Quiver.Hom.unop_op, hc₂, ← unop_comp_assoc, ← unop_comp_assoc, - ← opShiftFunctorEquivalence_unitIso_inv_naturality] - simp - -/-- The pretriangulated structure on the opposite category of -a pretriangulated category. It is a scoped instance, so that we need to -`open CategoryTheory.Pretriangulated.Opposite` in order to be able -to use it: the reason is that it relies on the definition of the shift -on the opposite category `Cᵒᵖ`, for which it is unclear whether it should -be a global instance or not. -/ -scoped instance : Pretriangulated Cᵒᵖ where - distinguishedTriangles := distinguishedTriangles C - isomorphic_distinguished := isomorphic_distinguished - contractible_distinguished := contractible_distinguished - distinguished_cocone_triangle := distinguished_cocone_triangle - rotate_distinguished_triangle := rotate_distinguished_triangle - complete_distinguished_triangle_morphism := complete_distinguished_triangle_morphism - -end Opposite - -variable {C} - -lemma mem_distTriang_op_iff (T : Triangle Cᵒᵖ) : - (T ∈ distTriang Cᵒᵖ) ↔ ((triangleOpEquivalence C).inverse.obj T).unop ∈ distTriang C := by - rfl - -lemma mem_distTriang_op_iff' (T : Triangle Cᵒᵖ) : - (T ∈ distTriang Cᵒᵖ) ↔ ∃ (T' : Triangle C) (_ : T' ∈ distTriang C), - Nonempty (T ≅ (triangleOpEquivalence C).functor.obj (Opposite.op T')) := - Opposite.mem_distinguishedTriangles_iff' T - -lemma op_distinguished (T : Triangle C) (hT : T ∈ distTriang C) : - ((triangleOpEquivalence C).functor.obj (Opposite.op T)) ∈ distTriang Cᵒᵖ := by - rw [mem_distTriang_op_iff'] - exact ⟨T, hT, ⟨Iso.refl _⟩⟩ - -lemma unop_distinguished (T : Triangle Cᵒᵖ) (hT : T ∈ distTriang Cᵒᵖ) : - ((triangleOpEquivalence C).inverse.obj T).unop ∈ distTriang C := hT - end Pretriangulated -namespace Functor - -open Pretriangulated.Opposite Pretriangulated - -variable {C} - -lemma map_distinguished_op_exact [HasShift C ℤ] [HasZeroObject C] [Preadditive C] - [∀ (n : ℤ), (shiftFunctor C n).Additive] - [Pretriangulated C]{A : Type*} [Category A] [Abelian A] (F : Cᵒᵖ ⥤ A) - [F.IsHomological] (T : Triangle C) (hT : T ∈ distTriang C) : - ((shortComplexOfDistTriangle T hT).op.map F).Exact := - F.map_distinguished_exact _ (op_distinguished T hT) - -end Functor - end CategoryTheory diff --git a/Mathlib/CategoryTheory/Triangulated/Opposite/Pretriangulated.lean b/Mathlib/CategoryTheory/Triangulated/Opposite/Pretriangulated.lean new file mode 100644 index 0000000000000..5e146b7317e28 --- /dev/null +++ b/Mathlib/CategoryTheory/Triangulated/Opposite/Pretriangulated.lean @@ -0,0 +1,195 @@ +/- +Copyright (c) 2023 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.Triangulated.Opposite.Triangle +import Mathlib.CategoryTheory.Triangulated.HomologicalFunctor + +/-! +# The (pre)triangulated structure on the opposite category + +In this file, we shall construct the (pre)triangulated structure +on the opposite category `Cᵒᵖ` of a (pre)triangulated category `C`. + +The shift on `Cᵒᵖ` was constructed in ``CategoryTheory.Triangulated.Opposite.Basic`, +and is such that shifting by `n : ℤ` on `Cᵒᵖ` corresponds to the shift by +`-n` on `C`. In `CategoryTheory.Triangulated.Opposite.Triangle`, we constructed +an equivalence `(Triangle C)ᵒᵖ ≌ Triangle Cᵒᵖ`, called +`CategoryTheory.Pretriangulated.triangleOpEquivalence`. + +Here, we defined the notion of distinguished triangles in `Cᵒᵖ`, such that +`triangleOpEquivalence` sends distinguished triangles in `C` to distinguished triangles +in `Cᵒᵖ`. In other words, if `X ⟶ Y ⟶ Z ⟶ X⟦1⟧` is a distinguished triangle in `C`, +then the triangle `op Z ⟶ op Y ⟶ op X ⟶ (op Z)⟦1⟧` that is deduced *without introducing signs* +shall be a distinguished triangle in `Cᵒᵖ`. This is equivalent to the definition +in [Verdiers's thesis, p. 96][verdier1996] which would require that the triangle +`(op X)⟦-1⟧ ⟶ op Z ⟶ op Y ⟶ op X` (without signs) is *antidistinguished*. + +## References +* [Jean-Louis Verdier, *Des catégories dérivées des catégories abéliennes*][verdier1996] + +-/ + +namespace CategoryTheory + +open Category Limits Preadditive ZeroObject + +variable (C : Type*) [Category C] [HasShift C ℤ] [HasZeroObject C] [Preadditive C] + [∀ (n : ℤ), (shiftFunctor C n).Additive] [Pretriangulated C] + +namespace Pretriangulated + +open Opposite + +namespace Opposite + +/-- A triangle in `Cᵒᵖ` shall be distinguished iff it corresponds to a distinguished +triangle in `C` via the equivalence `triangleOpEquivalence C : (Triangle C)ᵒᵖ ≌ Triangle Cᵒᵖ`. -/ +def distinguishedTriangles : Set (Triangle Cᵒᵖ) := + fun T => ((triangleOpEquivalence C).inverse.obj T).unop ∈ distTriang C + +variable {C} + +lemma mem_distinguishedTriangles_iff (T : Triangle Cᵒᵖ) : + T ∈ distinguishedTriangles C ↔ + ((triangleOpEquivalence C).inverse.obj T).unop ∈ distTriang C := by + rfl + +lemma mem_distinguishedTriangles_iff' (T : Triangle Cᵒᵖ) : + T ∈ distinguishedTriangles C ↔ + ∃ (T' : Triangle C) (_ : T' ∈ distTriang C), + Nonempty (T ≅ (triangleOpEquivalence C).functor.obj (Opposite.op T')) := by + rw [mem_distinguishedTriangles_iff] + constructor + · intro hT + exact ⟨_ ,hT, ⟨(triangleOpEquivalence C).counitIso.symm.app T⟩⟩ + · rintro ⟨T', hT', ⟨e⟩⟩ + refine isomorphic_distinguished _ hT' _ ?_ + exact Iso.unop ((triangleOpEquivalence C).unitIso.app (Opposite.op T') ≪≫ + (triangleOpEquivalence C).inverse.mapIso e.symm) + +lemma isomorphic_distinguished (T₁ : Triangle Cᵒᵖ) + (hT₁ : T₁ ∈ distinguishedTriangles C) (T₂ : Triangle Cᵒᵖ) (e : T₂ ≅ T₁) : + T₂ ∈ distinguishedTriangles C := by + simp only [mem_distinguishedTriangles_iff] at hT₁ ⊢ + exact Pretriangulated.isomorphic_distinguished _ hT₁ _ + ((triangleOpEquivalence C).inverse.mapIso e).unop.symm + +/-- Up to rotation, the contractible triangle `X ⟶ X ⟶ 0 ⟶ X⟦1⟧` for `X : Cᵒᵖ` corresponds +to the contractible triangle for `X.unop` in `C`. -/ +@[simps!] +noncomputable def contractibleTriangleIso (X : Cᵒᵖ) : + contractibleTriangle X ≅ (triangleOpEquivalence C).functor.obj + (Opposite.op (contractibleTriangle X.unop).invRotate) := + Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) + (IsZero.iso (isZero_zero _) (by + dsimp + rw [IsZero.iff_id_eq_zero] + change (𝟙 ((0 : C)⟦(-1 : ℤ)⟧)).op = 0 + rw [← Functor.map_id, id_zero, Functor.map_zero, op_zero])) + (by aesop_cat) (by aesop_cat) (by aesop_cat) + +lemma contractible_distinguished (X : Cᵒᵖ) : + contractibleTriangle X ∈ distinguishedTriangles C := by + rw [mem_distinguishedTriangles_iff'] + exact ⟨_, inv_rot_of_distTriang _ (Pretriangulated.contractible_distinguished X.unop), + ⟨contractibleTriangleIso X⟩⟩ + +/-- Isomorphism expressing a compatibility of the equivalence `triangleOpEquivalence C` +with the rotation of triangles. -/ +noncomputable def rotateTriangleOpEquivalenceInverseObjRotateUnopIso (T : Triangle Cᵒᵖ) : + ((triangleOpEquivalence C).inverse.obj T.rotate).unop.rotate ≅ + ((triangleOpEquivalence C).inverse.obj T).unop := + Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) + (-((opShiftFunctorEquivalence C 1).unitIso.app T.obj₁).unop) (by simp) + (Quiver.Hom.op_inj (by aesop_cat)) (by aesop_cat) + +lemma rotate_distinguished_triangle (T : Triangle Cᵒᵖ) : + T ∈ distinguishedTriangles C ↔ T.rotate ∈ distinguishedTriangles C := by + simp only [mem_distinguishedTriangles_iff, Pretriangulated.rotate_distinguished_triangle + ((triangleOpEquivalence C).inverse.obj (T.rotate)).unop] + exact distinguished_iff_of_iso (rotateTriangleOpEquivalenceInverseObjRotateUnopIso T).symm + +lemma distinguished_cocone_triangle {X Y : Cᵒᵖ} (f : X ⟶ Y) : + ∃ (Z : Cᵒᵖ) (g : Y ⟶ Z) (h : Z ⟶ X⟦(1 : ℤ)⟧), + Triangle.mk f g h ∈ distinguishedTriangles C := by + obtain ⟨Z, g, h, H⟩ := Pretriangulated.distinguished_cocone_triangle₁ f.unop + refine ⟨_, g.op, (opShiftFunctorEquivalence C 1).counitIso.inv.app (Opposite.op Z) ≫ + (shiftFunctor Cᵒᵖ (1 : ℤ)).map h.op, ?_⟩ + simp only [mem_distinguishedTriangles_iff] + refine Pretriangulated.isomorphic_distinguished _ H _ ?_ + exact Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) (by aesop_cat) (by aesop_cat) + (Quiver.Hom.op_inj (by simp [shift_unop_opShiftFunctorEquivalence_counitIso_inv_app])) + +lemma complete_distinguished_triangle_morphism (T₁ T₂ : Triangle Cᵒᵖ) + (hT₁ : T₁ ∈ distinguishedTriangles C) (hT₂ : T₂ ∈ distinguishedTriangles C) + (a : T₁.obj₁ ⟶ T₂.obj₁) (b : T₁.obj₂ ⟶ T₂.obj₂) (comm : T₁.mor₁ ≫ b = a ≫ T₂.mor₁) : + ∃ (c : T₁.obj₃ ⟶ T₂.obj₃), T₁.mor₂ ≫ c = b ≫ T₂.mor₂ ∧ + T₁.mor₃ ≫ a⟦1⟧' = c ≫ T₂.mor₃ := by + rw [mem_distinguishedTriangles_iff] at hT₁ hT₂ + obtain ⟨c, hc₁, hc₂⟩ := + Pretriangulated.complete_distinguished_triangle_morphism₁ _ _ hT₂ hT₁ + b.unop a.unop (Quiver.Hom.op_inj comm.symm) + dsimp at c hc₁ hc₂ + replace hc₂ := ((opShiftFunctorEquivalence C 1).unitIso.hom.app T₂.obj₁).unop ≫= hc₂ + dsimp at hc₂ + simp only [assoc, Iso.unop_hom_inv_id_app_assoc] at hc₂ + refine ⟨c.op, Quiver.Hom.unop_inj hc₁.symm, Quiver.Hom.unop_inj ?_⟩ + apply (shiftFunctor C (1 : ℤ)).map_injective + rw [unop_comp, unop_comp, Functor.map_comp, Functor.map_comp, + Quiver.Hom.unop_op, hc₂, ← unop_comp_assoc, ← unop_comp_assoc, + ← opShiftFunctorEquivalence_unitIso_inv_naturality] + simp + +/-- The pretriangulated structure on the opposite category of +a pretriangulated category. It is a scoped instance, so that we need to +`open CategoryTheory.Pretriangulated.Opposite` in order to be able +to use it: the reason is that it relies on the definition of the shift +on the opposite category `Cᵒᵖ`, for which it is unclear whether it should +be a global instance or not. -/ +scoped instance : Pretriangulated Cᵒᵖ where + distinguishedTriangles := distinguishedTriangles C + isomorphic_distinguished := isomorphic_distinguished + contractible_distinguished := contractible_distinguished + distinguished_cocone_triangle := distinguished_cocone_triangle + rotate_distinguished_triangle := rotate_distinguished_triangle + complete_distinguished_triangle_morphism := complete_distinguished_triangle_morphism + +end Opposite + +variable {C} + +lemma mem_distTriang_op_iff (T : Triangle Cᵒᵖ) : + (T ∈ distTriang Cᵒᵖ) ↔ ((triangleOpEquivalence C).inverse.obj T).unop ∈ distTriang C := by + rfl + +lemma mem_distTriang_op_iff' (T : Triangle Cᵒᵖ) : + (T ∈ distTriang Cᵒᵖ) ↔ ∃ (T' : Triangle C) (_ : T' ∈ distTriang C), + Nonempty (T ≅ (triangleOpEquivalence C).functor.obj (Opposite.op T')) := + Opposite.mem_distinguishedTriangles_iff' T + +lemma op_distinguished (T : Triangle C) (hT : T ∈ distTriang C) : + ((triangleOpEquivalence C).functor.obj (Opposite.op T)) ∈ distTriang Cᵒᵖ := by + rw [mem_distTriang_op_iff'] + exact ⟨T, hT, ⟨Iso.refl _⟩⟩ + +lemma unop_distinguished (T : Triangle Cᵒᵖ) (hT : T ∈ distTriang Cᵒᵖ) : + ((triangleOpEquivalence C).inverse.obj T).unop ∈ distTriang C := hT + +end Pretriangulated + +namespace Functor + +open Pretriangulated.Opposite Pretriangulated + +variable {C} + +lemma map_distinguished_op_exact {A : Type*} [Category A] [Abelian A] (F : Cᵒᵖ ⥤ A) + [F.IsHomological] (T : Triangle C) (hT : T ∈ distTriang C) : + ((shortComplexOfDistTriangle T hT).op.map F).Exact := + F.map_distinguished_exact _ (op_distinguished T hT) + +end Functor + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Triangulated/Opposite/Triangle.lean b/Mathlib/CategoryTheory/Triangulated/Opposite/Triangle.lean new file mode 100644 index 0000000000000..cd0a07b4be3dd --- /dev/null +++ b/Mathlib/CategoryTheory/Triangulated/Opposite/Triangle.lean @@ -0,0 +1,110 @@ +/- +Copyright (c) 2023 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.Triangulated.Basic +import Mathlib.CategoryTheory.Triangulated.Opposite.Basic + +/-! +# Triangles in the opposite category of a (pre)triangulated category + +Let `C` be a (pre)triangulated category. +In `CategoryTheory.Triangulated.Opposite.Basic`, we have constructed +a shift on `Cᵒᵖ` that will be part of a structure of (pre)triangulated +category. In this file, we construct an equivalence of categories +between `(Triangle C)ᵒᵖ` and `Triangle Cᵒᵖ`, called +`CategoryTheory.Pretriangulated.triangleOpEquivalence`. It sends a triangle +`X ⟶ Y ⟶ Z ⟶ X⟦1⟧` in `C` to the triangle `op Z ⟶ op Y ⟶ op X ⟶ (op Z)⟦1⟧` in `Cᵒᵖ` +(without introducing signs). + +## References +* [Jean-Louis Verdier, *Des catégories dérivées des catégories abéliennes*][verdier1996] + +-/ + +namespace CategoryTheory.Pretriangulated + +open Category Limits Preadditive ZeroObject Opposite + +variable (C : Type*) [Category C] [HasShift C ℤ] + +namespace TriangleOpEquivalence + +/-- The functor which sends a triangle `X ⟶ Y ⟶ Z ⟶ X⟦1⟧` in `C` to the triangle +`op Z ⟶ op Y ⟶ op X ⟶ (op Z)⟦1⟧` in `Cᵒᵖ` (without introducing signs). -/ +@[simps] +noncomputable def functor : (Triangle C)ᵒᵖ ⥤ Triangle Cᵒᵖ where + obj T := Triangle.mk T.unop.mor₂.op T.unop.mor₁.op + ((opShiftFunctorEquivalence C 1).counitIso.inv.app (Opposite.op T.unop.obj₁) ≫ + T.unop.mor₃.op⟦(1 : ℤ)⟧') + map {T₁ T₂} φ := + { hom₁ := φ.unop.hom₃.op + hom₂ := φ.unop.hom₂.op + hom₃ := φ.unop.hom₁.op + comm₁ := Quiver.Hom.unop_inj φ.unop.comm₂.symm + comm₂ := Quiver.Hom.unop_inj φ.unop.comm₁.symm + comm₃ := by + dsimp + rw [assoc, ← Functor.map_comp, ← op_comp, ← φ.unop.comm₃, op_comp, Functor.map_comp, + opShiftFunctorEquivalence_counitIso_inv_naturality_assoc] + rfl } + +/-- The functor which sends a triangle `X ⟶ Y ⟶ Z ⟶ X⟦1⟧` in `Cᵒᵖ` to the triangle +`Z.unop ⟶ Y.unop ⟶ X.unop ⟶ Z.unop⟦1⟧` in `C` (without introducing signs). -/ +@[simps] +noncomputable def inverse : Triangle Cᵒᵖ ⥤ (Triangle C)ᵒᵖ where + obj T := Opposite.op (Triangle.mk T.mor₂.unop T.mor₁.unop + (((opShiftFunctorEquivalence C 1).unitIso.inv.app T.obj₁).unop ≫ T.mor₃.unop⟦(1 : ℤ)⟧')) + map {T₁ T₂} φ := Quiver.Hom.op + { hom₁ := φ.hom₃.unop + hom₂ := φ.hom₂.unop + hom₃ := φ.hom₁.unop + comm₁ := Quiver.Hom.op_inj φ.comm₂.symm + comm₂ := Quiver.Hom.op_inj φ.comm₁.symm + comm₃ := Quiver.Hom.op_inj (by + dsimp + rw [assoc, ← opShiftFunctorEquivalence_unitIso_inv_naturality, + ← op_comp_assoc, ← Functor.map_comp, ← unop_comp, ← φ.comm₃, + unop_comp, Functor.map_comp, op_comp, assoc]) } + +/-- The unit isomorphism of the +equivalence `triangleOpEquivalence C : (Triangle C)ᵒᵖ ≌ Triangle Cᵒᵖ` . -/ +@[simps!] +noncomputable def unitIso : 𝟭 _ ≅ functor C ⋙ inverse C := + NatIso.ofComponents (fun T => Iso.op + (Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) (by aesop_cat) (by aesop_cat) + (Quiver.Hom.op_inj + (by simp [shift_unop_opShiftFunctorEquivalence_counitIso_inv_app])))) + (fun {T₁ T₂} f => Quiver.Hom.unop_inj (by aesop_cat)) + +/-- The counit isomorphism of the +equivalence `triangleOpEquivalence C : (Triangle C)ᵒᵖ ≌ Triangle Cᵒᵖ` . -/ +@[simps!] +noncomputable def counitIso : inverse C ⋙ functor C ≅ 𝟭 _ := + NatIso.ofComponents (fun T => by + refine Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) ?_ ?_ ?_ + · aesop_cat + · aesop_cat + · dsimp + rw [Functor.map_id, comp_id, id_comp, Functor.map_comp, + ← opShiftFunctorEquivalence_counitIso_inv_naturality_assoc, + opShiftFunctorEquivalence_counitIso_inv_app_shift, ← Functor.map_comp, + Iso.hom_inv_id_app, Functor.map_id] + simp only [Functor.id_obj, comp_id]) + (by aesop_cat) + +end TriangleOpEquivalence + +/-- An anti-equivalence between the categories of triangles in `C` and in `Cᵒᵖ`. +A triangle in `Cᵒᵖ` shall be distinguished iff it correspond to a distinguished +triangle in `C` via this equivalence. -/ +@[simps] +noncomputable def triangleOpEquivalence : + (Triangle C)ᵒᵖ ≌ Triangle Cᵒᵖ where + functor := TriangleOpEquivalence.functor C + inverse := TriangleOpEquivalence.inverse C + unitIso := TriangleOpEquivalence.unitIso C + counitIso := TriangleOpEquivalence.counitIso C + +end CategoryTheory.Pretriangulated diff --git a/Mathlib/CategoryTheory/Triangulated/Yoneda.lean b/Mathlib/CategoryTheory/Triangulated/Yoneda.lean index e136effd690d8..40fe591bc97b9 100644 --- a/Mathlib/CategoryTheory/Triangulated/Yoneda.lean +++ b/Mathlib/CategoryTheory/Triangulated/Yoneda.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.Homology.ShortComplex.Ab import Mathlib.CategoryTheory.Preadditive.Yoneda.Basic import Mathlib.CategoryTheory.Shift.ShiftedHomOpposite import Mathlib.CategoryTheory.Triangulated.HomologicalFunctor -import Mathlib.CategoryTheory.Triangulated.Opposite +import Mathlib.CategoryTheory.Triangulated.Opposite.Pretriangulated /-! # The Yoneda functors are homological From 68fbd252aa5d0168c5f582c0f66b3005dba683e1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 29 Dec 2024 12:22:20 +0000 Subject: [PATCH 029/189] feat: `(-1 : R) ^ n = ite (Even n) 1 (-1)` (#20320) Co-authored-by: Heather Macbeth --- Mathlib/Algebra/Ring/Parity.lean | 10 +++++++--- 1 file changed, 7 insertions(+), 3 deletions(-) diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index f74ff162454d5..91d436c897684 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -351,10 +351,14 @@ lemma iterate_eq_id (hf : Involutive f) (hne : f ≠ id) : f^[n] = id ↔ Even n end Involutive end Function +lemma neg_one_pow_eq_ite {R : Type*} [Monoid R] [HasDistribNeg R] {n : ℕ} : + (-1 : R) ^ n = ite (Even n) 1 (-1) := by + cases even_or_odd n with + | inl h => rw [h.neg_one_pow, if_pos h] + | inr h => rw [h.neg_one_pow, if_neg (by simpa using h)] + lemma neg_one_pow_eq_one_iff_even {R : Type*} [Monoid R] [HasDistribNeg R] {n : ℕ} - (h : (-1 : R) ≠ 1) : (-1 : R) ^ n = 1 ↔ Even n where - mp h' := of_not_not fun hn ↦ h <| (not_even_iff_odd.1 hn).neg_one_pow.symm.trans h' - mpr := Even.neg_one_pow + (h : (-1 : R) ≠ 1) : (-1 : R) ^ n = 1 ↔ Even n := by simp [neg_one_pow_eq_ite, h] section CharTwo From de084eb0b0314561f9dcfc53be76298d2e03b22c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 29 Dec 2024 13:25:55 +0000 Subject: [PATCH 030/189] chore(Data/List/MinMax): fix lemma names (#20318) --- Mathlib/Data/List/MinMax.lean | 45 ++++++++++++++++++--------------- scripts/nolints_prime_decls.txt | 2 +- 2 files changed, 26 insertions(+), 21 deletions(-) diff --git a/Mathlib/Data/List/MinMax.lean b/Mathlib/Data/List/MinMax.lean index 5ed4618178ea5..cb9db4834736d 100644 --- a/Mathlib/Data/List/MinMax.lean +++ b/Mathlib/Data/List/MinMax.lean @@ -291,19 +291,23 @@ theorem minimum_eq_top {l : List α} : l.minimum = ⊤ ↔ l = [] := @[simp, deprecated minimum_eq_top "Don't mix Option and WithTop" (since := "2024-05-27")] theorem minimum_eq_none {l : List α} : l.minimum = none ↔ l = [] := minimum_eq_top -theorem not_lt_maximum_of_mem : a ∈ l → (maximum l : WithBot α) = m → ¬m < a := +theorem not_maximum_lt_of_mem : a ∈ l → (maximum l : WithBot α) = m → ¬m < a := not_lt_of_mem_argmax -theorem minimum_not_lt_of_mem : a ∈ l → (minimum l : WithTop α) = m → ¬a < m := +@[deprecated (since := "2024-12-29")] alias not_lt_maximum_of_mem := not_maximum_lt_of_mem + +theorem not_lt_minimum_of_mem : a ∈ l → (minimum l : WithTop α) = m → ¬a < m := not_lt_of_mem_argmin -theorem not_lt_maximum_of_mem' (ha : a ∈ l) : ¬maximum l < (a : WithBot α) := by - cases h : l.maximum - · simp_all - · simp [not_lt_maximum_of_mem ha h, not_false_iff] +@[deprecated (since := "2024-12-29")] alias minimum_not_lt_of_mem := not_lt_minimum_of_mem + +theorem not_maximum_lt_of_mem' (ha : a ∈ l) : ¬maximum l < (a : WithBot α) := by + cases h : l.maximum <;> simp_all [not_maximum_lt_of_mem ha] -theorem not_lt_minimum_of_mem' (ha : a ∈ l) : ¬(a : WithTop α) < minimum l := - @not_lt_maximum_of_mem' αᵒᵈ _ _ _ _ ha +@[deprecated (since := "2024-12-29")] alias not_lt_maximum_of_mem' := not_maximum_lt_of_mem' + +theorem not_lt_minimum_of_mem' (ha : a ∈ l) : ¬(a : WithTop α) < minimum l := by + cases h : l.minimum <;> simp_all [not_lt_minimum_of_mem ha] end Preorder @@ -324,10 +328,10 @@ theorem minimum_le_of_mem : a ∈ l → (minimum l : WithTop α) = m → m ≤ a le_of_mem_argmin theorem le_maximum_of_mem' (ha : a ∈ l) : (a : WithBot α) ≤ maximum l := - le_of_not_lt <| not_lt_maximum_of_mem' ha + le_of_not_lt <| not_maximum_lt_of_mem' ha theorem minimum_le_of_mem' (ha : a ∈ l) : minimum l ≤ (a : WithTop α) := - @le_maximum_of_mem' αᵒᵈ _ _ _ ha + le_of_not_lt <| not_lt_minimum_of_mem' ha theorem minimum_concat (a : α) (l : List α) : minimum (l ++ [a]) = min (minimum l) a := @maximum_concat αᵒᵈ _ _ _ @@ -343,11 +347,15 @@ theorem maximum_le_of_forall_le {b : WithBot α} (h : ∀ a ∈ l, a ≤ b) : l. induction l with | nil => simp | cons a l ih => - simp only [maximum_cons, max_le_iff, WithBot.coe_le_coe] + simp only [maximum_cons, max_le_iff] exact ⟨h a (by simp), ih fun a w => h a (mem_cons.mpr (Or.inr w))⟩ -theorem le_minimum_of_forall_le {b : WithTop α} (h : ∀ a ∈ l, b ≤ a) : b ≤ l.minimum := - maximum_le_of_forall_le (α := αᵒᵈ) h +theorem le_minimum_of_forall_le {b : WithTop α} (h : ∀ a ∈ l, b ≤ a) : b ≤ l.minimum := by + induction l with + | nil => simp + | cons a l ih => + simp only [minimum_cons, le_min_iff] + exact ⟨h a (by simp), ih fun a w => h a (mem_cons.mpr (Or.inr w))⟩ theorem maximum_eq_coe_iff : maximum l = m ↔ m ∈ l ∧ ∀ a ∈ l, a ≤ m := by rw [maximum, ← WithBot.some_eq_coe, argmax_eq_some_iff] @@ -359,13 +367,10 @@ theorem minimum_eq_coe_iff : minimum l = m ↔ m ∈ l ∧ ∀ a ∈ l, m ≤ a @maximum_eq_coe_iff αᵒᵈ _ _ _ theorem coe_le_maximum_iff : a ≤ l.maximum ↔ ∃ b, b ∈ l ∧ a ≤ b := by - induction l with - | nil => simp - | cons h t ih => - simp [maximum_cons, ih] + induction' l <;> simp [maximum_cons, *] -theorem minimum_le_coe_iff : l.minimum ≤ a ↔ ∃ b, b ∈ l ∧ b ≤ a := - coe_le_maximum_iff (α := αᵒᵈ) +theorem minimum_le_coe_iff : l.minimum ≤ a ↔ ∃ b, b ∈ l ∧ b ≤ a := by + induction' l <;> simp [minimum_cons, *] theorem maximum_ne_bot_of_ne_nil (h : l ≠ []) : l.maximum ≠ ⊥ := match l, h with | _ :: _, _ => by simp [maximum_cons] @@ -405,7 +410,7 @@ theorem le_maximum_of_length_pos_iff {b : α} (h : 0 < l.length) : @[simp] theorem minimum_of_length_pos_le_iff {b : α} (h : 0 < l.length) : minimum_of_length_pos h ≤ b ↔ l.minimum ≤ b := - le_maximum_of_length_pos_iff (α := αᵒᵈ) h + WithTop.untop_le_iff _ theorem maximum_of_length_pos_mem (h : 0 < l.length) : maximum_of_length_pos h ∈ l := by diff --git a/scripts/nolints_prime_decls.txt b/scripts/nolints_prime_decls.txt index 47c01ac4b1533..253e549ada88b 100644 --- a/scripts/nolints_prime_decls.txt +++ b/scripts/nolints_prime_decls.txt @@ -2417,7 +2417,7 @@ List.next_cons_cons_eq' List.nnnorm_prod_le' List.nodup_sublists' List.norm_prod_le' -List.not_lt_maximum_of_mem' +List.not_maximum_lt_of_mem' List.not_lt_minimum_of_mem' List.ofFn_succ' List.Pairwise.chain' From 58dc043d6bdac239bfbdcb8de5257893aac92baa Mon Sep 17 00:00:00 2001 From: LessnessRandomness Date: Sun, 29 Dec 2024 17:38:54 +0000 Subject: [PATCH 031/189] docs(Combinatorics/SimpleGraph): typo in Maps.lean (#20327) Change `G` to `G'` in the documentation of graph embedding. From b4e1f959102f22b2826be398c7af96518e8124a6 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sun, 29 Dec 2024 20:14:48 +0000 Subject: [PATCH 032/189] feat(Order/KrullDimension): equivalent conditions for zero-dimensional, one-dimensional etc (#20176) --- Mathlib/Data/ENat/Basic.lean | 14 ++ Mathlib/Order/KrullDimension.lean | 138 ++++++++++++++----- Mathlib/Order/RelSeries.lean | 43 ++++++ Mathlib/RingTheory/KrullDimension/Basic.lean | 4 +- 4 files changed, 165 insertions(+), 34 deletions(-) diff --git a/Mathlib/Data/ENat/Basic.lean b/Mathlib/Data/ENat/Basic.lean index e5b176f751255..7f2011f33653d 100644 --- a/Mathlib/Data/ENat/Basic.lean +++ b/Mathlib/Data/ENat/Basic.lean @@ -461,3 +461,17 @@ protected def _root_.RingHom.ENatMap {S : Type*} [CanonicallyOrderedCommSemiring {MonoidWithZeroHom.ENatMap f.toMonoidWithZeroHom hf, f.toAddMonoidHom.ENatMap with} end ENat + +lemma WithBot.lt_add_one_iff {n : WithBot ℕ∞} {m : ℕ} : n < m + 1 ↔ n ≤ m := by + rw [← WithBot.coe_one, ← ENat.coe_one, WithBot.coe_natCast, ← Nat.cast_add, ← WithBot.coe_natCast] + cases n + · simp only [bot_le, iff_true, WithBot.bot_lt_coe] + · rw [WithBot.coe_lt_coe, Nat.cast_add, ENat.coe_one, ENat.lt_add_one_iff (ENat.coe_ne_top _), + ← WithBot.coe_le_coe, WithBot.coe_natCast] + +lemma WithBot.add_one_le_iff {n : ℕ} {m : WithBot ℕ∞} : n + 1 ≤ m ↔ n < m := by + rw [← WithBot.coe_one, ← ENat.coe_one, WithBot.coe_natCast, ← Nat.cast_add, ← WithBot.coe_natCast] + cases m + · simp + · rw [WithBot.coe_le_coe, ENat.coe_add, ENat.coe_one, ENat.add_one_le_iff (ENat.coe_ne_top n), + ← WithBot.coe_lt_coe, WithBot.coe_natCast] diff --git a/Mathlib/Order/KrullDimension.lean b/Mathlib/Order/KrullDimension.lean index 3a97b0b3185a7..f9939bd2e168b 100644 --- a/Mathlib/Order/KrullDimension.lean +++ b/Mathlib/Order/KrullDimension.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Order.Group.Int import Mathlib.Data.ENat.Lattice import Mathlib.Order.Minimal import Mathlib.Order.RelSeries +import Mathlib.Tactic.FinCases /-! # Krull dimension of a preordered set and height of an element @@ -508,19 +509,77 @@ variable [Preorder α] [Preorder β] lemma LTSeries.length_le_krullDim (p : LTSeries α) : p.length ≤ krullDim α := le_sSup ⟨_, rfl⟩ -lemma krullDim_nonneg_of_nonempty [Nonempty α] : 0 ≤ krullDim α := - le_sSup ⟨⟨0, fun _ ↦ @Nonempty.some α inferInstance, fun f ↦ f.elim0⟩, rfl⟩ +lemma krullDim_eq_bot_iff : krullDim α = ⊥ ↔ IsEmpty α := by + rw [eq_bot_iff, krullDim, iSup_le_iff] + simp only [le_bot_iff, WithBot.natCast_ne_bot, isEmpty_iff] + exact ⟨fun H x ↦ H ⟨0, fun _ ↦ x, by simp⟩, (· <| · 1)⟩ -/-- A definition of krullDim for nonempty `α` that avoids `WithBot` -/ -lemma krullDim_eq_iSup_length [Nonempty α] : - krullDim α = ⨆ (p : LTSeries α), (p.length : ℕ∞) := by - unfold krullDim - rw [WithBot.coe_iSup (OrderTop.bddAbove _)] - rfl +lemma krullDim_nonneg_iff : 0 ≤ krullDim α ↔ Nonempty α := by + rw [← not_iff_not, not_le, not_nonempty_iff, ← krullDim_eq_bot_iff, ← WithBot.lt_coe_bot, + bot_eq_zero, WithBot.coe_zero] + +lemma krullDim_eq_bot [IsEmpty α] : krullDim α = ⊥ := krullDim_eq_bot_iff.mpr ‹_› + +@[deprecated (since := "2024-12-22")] alias krullDim_eq_bot_of_isEmpty := krullDim_eq_bot + +lemma krullDim_nonneg [Nonempty α] : 0 ≤ krullDim α := krullDim_nonneg_iff.mpr ‹_› + +@[deprecated (since := "2024-12-22")] alias krullDim_nonneg_of_nonempty := krullDim_nonneg + +lemma krullDim_nonpos_iff_forall_isMax : krullDim α ≤ 0 ↔ ∀ x : α, IsMax x := by + simp only [krullDim, iSup_le_iff, isMax_iff_forall_not_lt] + refine ⟨fun H x y h ↦ (H ⟨1, ![x, y], + fun i ↦ by obtain rfl := Subsingleton.elim i 0; simpa⟩).not_lt (by simp), ?_⟩ + · rintro H ⟨_ | n, l, h⟩ + · simp + · cases H (l 0) (l 1) (h 0) -lemma krullDim_eq_bot_of_isEmpty [IsEmpty α] : krullDim α = ⊥ := WithBot.ciSup_empty _ +lemma krullDim_nonpos_iff_forall_isMin : krullDim α ≤ 0 ↔ ∀ x : α, IsMin x := by + simp only [krullDim_nonpos_iff_forall_isMax, IsMax, IsMin] + exact forall_swap -lemma krullDim_eq_top_of_infiniteDimensionalOrder [InfiniteDimensionalOrder α] : +lemma krullDim_le_one_iff : krullDim α ≤ 1 ↔ ∀ x : α, IsMin x ∨ IsMax x := by + rw [← not_iff_not] + simp_rw [isMax_iff_forall_not_lt, isMin_iff_forall_not_lt, krullDim, iSup_le_iff] + push_neg + constructor + · rintro ⟨⟨_ | _ | n, l, hl⟩, hl'⟩ + iterate 2 · cases hl'.not_le (by simp) + exact ⟨l 1, ⟨l 0, hl 0⟩, l 2, hl 1⟩ + · rintro ⟨x, ⟨y, hxy⟩, z, hzx⟩ + exact ⟨⟨2, ![y, x, z], fun i ↦ by fin_cases i <;> simpa⟩, by simp⟩ + +lemma krullDim_le_one_iff_forall_isMax {α : Type*} [PartialOrder α] [OrderBot α] : + krullDim α ≤ 1 ↔ ∀ x : α, x ≠ ⊥ → IsMax x := by + simp [krullDim_le_one_iff, ← or_iff_not_imp_left] + +lemma krullDim_le_one_iff_forall_isMin {α : Type*} [PartialOrder α] [OrderTop α] : + krullDim α ≤ 1 ↔ ∀ x : α, x ≠ ⊤ → IsMin x := by + simp [krullDim_le_one_iff, ← or_iff_not_imp_right] + +lemma krullDim_pos_iff : 0 < krullDim α ↔ ∃ x y : α, x < y := by + rw [← not_iff_not] + push_neg + simp_rw [← isMax_iff_forall_not_lt, ← krullDim_nonpos_iff_forall_isMax] + +lemma one_le_krullDim_iff : 1 ≤ krullDim α ↔ ∃ x y : α, x < y := by + rw [← krullDim_pos_iff, ← Nat.cast_zero, ← WithBot.add_one_le_iff, Nat.cast_zero, zero_add] + +lemma krullDim_nonpos_of_subsingleton [Subsingleton α] : krullDim α ≤ 0 := by + rw [krullDim_nonpos_iff_forall_isMax] + exact fun x y h ↦ (Subsingleton.elim x y).ge + +lemma krullDim_eq_zero_of_unique [Unique α] : krullDim α = 0 := + le_antisymm krullDim_nonpos_of_subsingleton krullDim_nonneg + +lemma krullDim_eq_length_of_finiteDimensionalOrder [FiniteDimensionalOrder α] : + krullDim α = (LTSeries.longestOf α).length := + le_antisymm + (iSup_le <| fun _ ↦ WithBot.coe_le_coe.mpr <| WithTop.coe_le_coe.mpr <| + RelSeries.length_le_length_longestOf _ _) <| + le_iSup (fun (i : LTSeries _) ↦ (i.length : WithBot (WithTop ℕ))) <| LTSeries.longestOf _ + +lemma krullDim_eq_top [InfiniteDimensionalOrder α] : krullDim α = ⊤ := le_antisymm le_top <| le_iSup_iff.mpr <| fun m hm ↦ match m, hm with | ⊥, hm => False.elim <| by @@ -532,29 +591,44 @@ lemma krullDim_eq_top_of_infiniteDimensionalOrder [InfiniteDimensionalOrder α] erw [WithBot.coe_lt_coe, WithTop.coe_lt_coe] simp -lemma krullDim_le_of_strictMono (f : α → β) (hf : StrictMono f) : krullDim α ≤ krullDim β := - iSup_le <| fun p ↦ le_sSup ⟨p.map f hf, rfl⟩ +@[deprecated (since := "2024-12-22")] +alias krullDim_eq_top_of_infiniteDimensionalOrder := krullDim_eq_top + +lemma krullDim_eq_top_iff : krullDim α = ⊤ ↔ InfiniteDimensionalOrder α := by + refine ⟨fun h ↦ ?_, fun _ ↦ krullDim_eq_top⟩ + cases isEmpty_or_nonempty α + · simp [krullDim_eq_bot] at h + cases finiteDimensionalOrder_or_infiniteDimensionalOrder α + · rw [krullDim_eq_length_of_finiteDimensionalOrder] at h + cases h + · infer_instance + +lemma le_krullDim_iff {n : ℕ} : n ≤ krullDim α ↔ ∃ l : LTSeries α, l.length = n := by + cases isEmpty_or_nonempty α + · simp [krullDim_eq_bot] + cases finiteDimensionalOrder_or_infiniteDimensionalOrder α + · rw [krullDim_eq_length_of_finiteDimensionalOrder, Nat.cast_le] + constructor + · exact fun H ↦ ⟨(LTSeries.longestOf α).take ⟨_, Nat.lt_succ.mpr H⟩, rfl⟩ + · exact fun ⟨l, hl⟩ ↦ hl ▸ l.longestOf_is_longest + · simpa [krullDim_eq_top] using Rel.InfiniteDimensional.exists_relSeries_with_length n -lemma krullDim_eq_length_of_finiteDimensionalOrder [FiniteDimensionalOrder α] : - krullDim α = (LTSeries.longestOf α).length := - le_antisymm - (iSup_le <| fun _ ↦ WithBot.coe_le_coe.mpr <| WithTop.coe_le_coe.mpr <| - RelSeries.length_le_length_longestOf _ _) <| - le_iSup (fun (i : LTSeries _) ↦ (i.length : WithBot (WithTop ℕ))) <| LTSeries.longestOf _ +/-- A definition of krullDim for nonempty `α` that avoids `WithBot` -/ +lemma krullDim_eq_iSup_length [Nonempty α] : + krullDim α = ⨆ (p : LTSeries α), (p.length : ℕ∞) := by + unfold krullDim + rw [WithBot.coe_iSup (OrderTop.bddAbove _)] + rfl -lemma krullDim_eq_zero_of_unique [Unique α] : krullDim α = 0 := by - rw [krullDim_eq_length_of_finiteDimensionalOrder (α := α), Nat.cast_eq_zero] - refine (LTSeries.longestOf_len_unique (default : LTSeries α) fun q ↦ show _ ≤ 0 from ?_).symm - by_contra r - exact ne_of_lt (q.step ⟨0, not_le.mp r⟩) <| Subsingleton.elim _ _ +lemma krullDim_lt_coe_iff {n : ℕ} : krullDim α < n ↔ ∀ l : LTSeries α, l.length < n := by + rw [krullDim, ← WithBot.coe_natCast] + cases' n with n + · rw [ENat.coe_zero, ← bot_eq_zero, WithBot.lt_coe_bot] + simp + · simp [WithBot.lt_add_one_iff, WithBot.coe_natCast, Nat.lt_succ] -lemma krullDim_nonpos_of_subsingleton [Subsingleton α] : krullDim α ≤ 0 := by - by_cases hα : Nonempty α - · have := uniqueOfSubsingleton (Classical.choice hα) - exact le_of_eq krullDim_eq_zero_of_unique - · have := not_nonempty_iff.mp hα - exact le_of_lt <| lt_of_eq_of_lt krullDim_eq_bot_of_isEmpty <| - Batteries.compareOfLessAndEq_eq_lt.mp rfl +lemma krullDim_le_of_strictMono (f : α → β) (hf : StrictMono f) : krullDim α ≤ krullDim β := + iSup_le fun p ↦ le_sSup ⟨p.map f hf, rfl⟩ lemma krullDim_le_of_strictComono_and_surj (f : α → β) (hf : ∀ ⦃a b⦄, f a < f b → a < b) (hf' : Function.Surjective f) : @@ -644,7 +718,7 @@ If `α` is `Nonempty`, then `krullDim_eq_iSup_height_of_nonempty`, with the coer -/ lemma krullDim_eq_iSup_height : krullDim α = ⨆ (a : α), ↑(height a) := by cases isEmpty_or_nonempty α with - | inl h => rw [krullDim_eq_bot_of_isEmpty, ciSup_of_empty] + | inl h => rw [krullDim_eq_bot, ciSup_of_empty] | inr h => rw [krullDim_eq_iSup_height_of_nonempty, WithBot.coe_iSup (OrderTop.bddAbove _)] /-- @@ -655,7 +729,7 @@ If `α` is `Nonempty`, then `krullDim_eq_iSup_coheight_of_nonempty`, with the co -/ lemma krullDim_eq_iSup_coheight : krullDim α = ⨆ (a : α), ↑(coheight a) := by cases isEmpty_or_nonempty α with - | inl h => rw [krullDim_eq_bot_of_isEmpty, ciSup_of_empty] + | inl h => rw [krullDim_eq_bot, ciSup_of_empty] | inr h => rw [krullDim_eq_iSup_coheight_of_nonempty, WithBot.coe_iSup (OrderTop.bddAbove _)] @[simp] -- not as useful as a simp lemma as it looks, due to the coe on the left diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index 3b0362616ff38..6f705bc25482f 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -8,6 +8,7 @@ import Mathlib.Data.Fintype.Card import Mathlib.Data.Fintype.Pi import Mathlib.Data.Fintype.Sigma import Mathlib.Data.Rel +import Mathlib.Data.Fin.VecNotation /-! # Series of a relation @@ -129,6 +130,7 @@ namespace Rel /-- A relation `r` is said to be finite dimensional iff there is a relation series of `r` with the maximum length. -/ +@[mk_iff] class FiniteDimensional : Prop where /-- A relation `r` is said to be finite dimensional iff there is a relation series of `r` with the maximum length. -/ @@ -136,6 +138,7 @@ class FiniteDimensional : Prop where /-- A relation `r` is said to be infinite dimensional iff there exists relation series of arbitrary length. -/ +@[mk_iff] class InfiniteDimensional : Prop where /-- A relation `r` is said to be infinite dimensional iff there exists relation series of arbitrary length. -/ @@ -620,6 +623,33 @@ lemma last_drop (p : RelSeries r) (i : Fin (p.length + 1)) : (p.drop i).last = p end RelSeries +variable {r} in +lemma Rel.not_finiteDimensional_iff [Nonempty α] : + ¬ r.FiniteDimensional ↔ r.InfiniteDimensional := by + rw [finiteDimensional_iff, infiniteDimensional_iff] + push_neg + constructor + · intro H n + induction n with + | zero => refine ⟨⟨0, ![Nonempty.some ‹_›], by simp⟩, by simp⟩ + | succ n IH => + obtain ⟨l, hl⟩ := IH + obtain ⟨l', hl'⟩ := H l + exact ⟨l'.take ⟨n + 1, by simpa [hl] using hl'⟩, rfl⟩ + · intro H l + obtain ⟨l', hl'⟩ := H (l.length + 1) + exact ⟨l', by simp [hl']⟩ + +variable {r} in +lemma Rel.not_infiniteDimensional_iff [Nonempty α] : + ¬ r.InfiniteDimensional ↔ r.FiniteDimensional := by + rw [← not_finiteDimensional_iff, not_not] + +lemma Rel.finiteDimensional_or_infiniteDimensional [Nonempty α] : + r.FiniteDimensional ∨ r.InfiniteDimensional := by + rw [← not_finiteDimensional_iff] + exact em r.FiniteDimensional + /-- A type is finite dimensional if its `LTSeries` has bounded length. -/ abbrev FiniteDimensionalOrder (γ : Type*) [Preorder γ] := Rel.FiniteDimensional ((· < ·) : γ → γ → Prop) @@ -815,6 +845,19 @@ end LTSeries end LTSeries +lemma not_finiteDimensionalOrder_iff [Preorder α] [Nonempty α] : + ¬ FiniteDimensionalOrder α ↔ InfiniteDimensionalOrder α := + Rel.not_finiteDimensional_iff + +lemma not_infiniteDimensionalOrder_iff [Preorder α] [Nonempty α] : + ¬ InfiniteDimensionalOrder α ↔ FiniteDimensionalOrder α := + Rel.not_infiniteDimensional_iff + +variable (α) in +lemma finiteDimensionalOrder_or_infiniteDimensionalOrder [Preorder α] [Nonempty α] : + FiniteDimensionalOrder α ∨ InfiniteDimensionalOrder α := + Rel.finiteDimensional_or_infiniteDimensional _ + /-- If `f : α → β` is a strictly monotonic function and `α` is an infinite dimensional type then so is `β`. -/ lemma infiniteDimensionalOrder_of_strictMono [Preorder α] [Preorder β] diff --git a/Mathlib/RingTheory/KrullDimension/Basic.lean b/Mathlib/RingTheory/KrullDimension/Basic.lean index cca9f211a88e5..a726c8a8a17d7 100644 --- a/Mathlib/RingTheory/KrullDimension/Basic.lean +++ b/Mathlib/RingTheory/KrullDimension/Basic.lean @@ -29,11 +29,11 @@ variable {R S : Type*} [CommRing R] [CommRing S] @[nontriviality] lemma ringKrullDim_eq_bot_of_subsingleton [Subsingleton R] : ringKrullDim R = ⊥ := - krullDim_eq_bot_of_isEmpty + krullDim_eq_bot lemma ringKrullDim_nonneg_of_nontrivial [Nontrivial R] : 0 ≤ ringKrullDim R := - krullDim_nonneg_of_nonempty + krullDim_nonneg /-- If `f : R →+* S` is surjective, then `ringKrullDim S ≤ ringKrullDim R`. -/ theorem ringKrullDim_le_of_surjective (f : R →+* S) (hf : Function.Surjective f) : From 3f703ecb9714a08b92a17f9c3c19ffc32343fede Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 29 Dec 2024 20:43:23 +0000 Subject: [PATCH 033/189] feat: quotienting out a polynomial in a family doesn't change the span (#20297) From GrowthInGroups (LeanCamCombi) Co-authored-by: Andrew Yang --- Mathlib/Algebra/Polynomial/CoeffMem.lean | 11 ++++++++- Mathlib/LinearAlgebra/Span/Defs.lean | 30 ++++++++++++++++++------ 2 files changed, 33 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/CoeffMem.lean b/Mathlib/Algebra/Polynomial/CoeffMem.lean index 1d6f5716ce825..118635b5bf511 100644 --- a/Mathlib/Algebra/Polynomial/CoeffMem.lean +++ b/Mathlib/Algebra/Polynomial/CoeffMem.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies, Andrew Yang -/ import Mathlib.Algebra.Algebra.Operations import Mathlib.Algebra.Polynomial.Div +import Mathlib.RingTheory.Ideal.Span /-! # Bounding the coefficients of the quotient and remainder of polynomials @@ -17,7 +18,7 @@ of `q`. -/ namespace Polynomial -variable {R S : Type*} [CommRing R] [Ring S] [Algebra R S] +variable {ι R S : Type*} [CommRing R] [Ring S] [Algebra R S] local notation "deg("p")" => natDegree p local notation3 "coeffs("p")" => Set.range (coeff p) @@ -101,4 +102,12 @@ lemma coeff_divByMonic_mem_pow_natDegree_mul (p q : S[X]) gcongr <;> exact sup_le (by simpa) (by simpa [Submodule.span_le, Set.range_subset_iff]) · simp +variable [DecidableEq ι] {i j : ι} + +open Function Ideal in +lemma idealSpan_range_update_divByMonic (hij : i ≠ j) (v : ι → R[X]) (hi : (v i).Monic) : + span (Set.range (Function.update v j (v j %ₘ v i))) = span (Set.range v) := by + rw [modByMonic_eq_sub_mul_div _ hi, mul_comm, ← smul_eq_mul, Ideal.span, Ideal.span, + Submodule.span_range_update_sub_smul hij] + end Polynomial diff --git a/Mathlib/LinearAlgebra/Span/Defs.lean b/Mathlib/LinearAlgebra/Span/Defs.lean index 2c9e703ae0c13..071461bc35a4f 100644 --- a/Mathlib/LinearAlgebra/Span/Defs.lean +++ b/Mathlib/LinearAlgebra/Span/Defs.lean @@ -358,12 +358,6 @@ theorem sup_toAddSubmonoid : (p ⊔ p').toAddSubmonoid = p.toAddSubmonoid ⊔ p' rw [mem_toAddSubmonoid, mem_sup, AddSubmonoid.mem_sup] rfl -theorem sup_toAddSubgroup {R M : Type*} [Ring R] [AddCommGroup M] [Module R M] - (p p' : Submodule R M) : (p ⊔ p').toAddSubgroup = p.toAddSubgroup ⊔ p'.toAddSubgroup := by - ext x - rw [mem_toAddSubgroup, mem_sup, AddSubgroup.mem_sup] - rfl - end theorem mem_span_singleton_self (x : M) : x ∈ R ∙ x := @@ -521,7 +515,13 @@ end AddCommMonoid section AddCommGroup -variable [Ring R] [AddCommGroup M] [Module R M] +variable [Ring R] [AddCommGroup M] [Module R M] {ι : Type*} [DecidableEq ι] {i j : ι} + +lemma sup_toAddSubgroup (p p' : Submodule R M) : + (p ⊔ p').toAddSubgroup = p.toAddSubgroup ⊔ p'.toAddSubgroup := by + ext x + rw [mem_toAddSubgroup, mem_sup, AddSubgroup.mem_sup] + rfl theorem mem_span_insert' {x y} {s : Set M} : x ∈ span R (insert y s) ↔ ∃ a : R, x + a • y ∈ span R s := by @@ -531,6 +531,22 @@ theorem mem_span_insert' {x y} {s : Set M} : · rintro ⟨a, h⟩ exact ⟨-a, _, h, by simp [add_comm, add_left_comm]⟩ +lemma span_range_update_add_smul (hij : i ≠ j) (v : ι → M) (r : R) : + span R (Set.range (Function.update v j (v j + r • v i))) = span R (Set.range v) := by + refine le_antisymm ?_ ?_ <;> simp only [span_le, Set.range_subset_iff, SetLike.mem_coe] <;> + intro k <;> obtain rfl | hjk := eq_or_ne j k + · rw [update_self] + exact add_mem (subset_span ⟨j, rfl⟩) <| smul_mem _ _ <| subset_span ⟨i, rfl⟩ + · exact subset_span ⟨k, (update_of_ne hjk.symm ..).symm⟩ + · nth_rw 2 [← add_sub_cancel_right (v j) (r • v i)] + exact sub_mem (subset_span ⟨j, update_self ..⟩) + (smul_mem _ _ (subset_span ⟨i, update_of_ne hij ..⟩)) + · exact subset_span ⟨k, update_of_ne hjk.symm ..⟩ + +lemma span_range_update_sub_smul (hij : i ≠ j) (v : ι → M) (r : R) : + span R (Set.range (Function.update v j (v j - r • v i))) = span R (Set.range v) := by + rw [sub_eq_add_neg, ← neg_smul, span_range_update_add_smul hij] + end AddCommGroup end Submodule From b4e6697e701bb49bcf49c504b8335747f55171c8 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Sun, 29 Dec 2024 22:09:48 +0000 Subject: [PATCH 034/189] chore: rename Nat.lt_div_iff_mul_lt (#20331) To make place for a lemma with this name in lean 4.16 --- Mathlib/Data/Nat/Defs.lean | 8 ++++++-- Mathlib/NumberTheory/Divisors.lean | 2 +- Mathlib/NumberTheory/FermatPsp.lean | 2 +- 3 files changed, 8 insertions(+), 4 deletions(-) diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index e1d2bb4909d02..1e19ccb7c114d 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -978,11 +978,15 @@ lemma mod_two_ne_zero : n % 2 ≠ 0 ↔ n % 2 = 1 := mod_two_not_eq_zero lemma div_mod_eq_mod_mul_div (a b c : ℕ) : a / b % c = a % (b * c) / b := (mod_mul_right_div_self a b c).symm -protected lemma lt_div_iff_mul_lt (hdn : d ∣ n) (a : ℕ) : a < n / d ↔ d * a < n := by +/-- Variant of `Nat.lt_div_iff_mul_lt` (added in Lean 4.16) that assumes `d ∣ n`. -/ +protected lemma lt_div_iff_mul_lt' (hdn : d ∣ n) (a : ℕ) : a < n / d ↔ d * a < n := by obtain rfl | hd := d.eq_zero_or_pos · simp [Nat.zero_dvd.1 hdn] · rw [← Nat.mul_lt_mul_left hd, ← Nat.eq_mul_of_div_eq_right hdn rfl] +@[deprecated Nat.lt_div_iff_mul_lt' (since := "2024-12-29")] +protected alias lt_div_iff_mul_lt := Nat.lt_div_iff_mul_lt' + lemma mul_div_eq_iff_dvd {n d : ℕ} : d * (n / d) = n ↔ d ∣ n := calc d * (n / d) = n ↔ d * (n / d) = d * (n / d) + (n % d) := by rw [div_add_mod] @@ -1291,7 +1295,7 @@ lemma dvd_left_injective : Function.Injective ((· ∣ ·) : ℕ → ℕ → Pro dvd_right_iff_eq.mp fun a => iff_of_eq (congr_fun h a) lemma div_lt_div_of_lt_of_dvd {a b d : ℕ} (hdb : d ∣ b) (h : a < b) : a / d < b / d := by - rw [Nat.lt_div_iff_mul_lt hdb] + rw [Nat.lt_div_iff_mul_lt' hdb] exact lt_of_le_of_lt (mul_div_le a d) h /-! ### Decidability of predicates -/ diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 5d34808cc4719..e3d2f24689a32 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -203,7 +203,7 @@ lemma one_lt_of_mem_properDivisors {m n : ℕ} (h : m ∈ n.properDivisors) : 1 lemma one_lt_div_of_mem_properDivisors {m n : ℕ} (h : m ∈ n.properDivisors) : 1 < n / m := by obtain ⟨h_dvd, h_lt⟩ := mem_properDivisors.mp h - rwa [Nat.lt_div_iff_mul_lt h_dvd, mul_one] + rwa [Nat.lt_div_iff_mul_lt' h_dvd, mul_one] /-- See also `Nat.mem_properDivisors`. -/ lemma mem_properDivisors_iff_exists {m n : ℕ} (hn : n ≠ 0) : diff --git a/Mathlib/NumberTheory/FermatPsp.lean b/Mathlib/NumberTheory/FermatPsp.lean index a620fe11fa855..12a27a0b7613d 100644 --- a/Mathlib/NumberTheory/FermatPsp.lean +++ b/Mathlib/NumberTheory/FermatPsp.lean @@ -126,7 +126,7 @@ section HelperLemmas private theorem a_id_helper {a b : ℕ} (ha : 2 ≤ a) (hb : 2 ≤ b) : 2 ≤ (a ^ b - 1) / (a - 1) := by change 1 < _ have h₁ : a - 1 ∣ a ^ b - 1 := by simpa only [one_pow] using nat_sub_dvd_pow_sub_pow a 1 b - rw [Nat.lt_div_iff_mul_lt h₁, mul_one, tsub_lt_tsub_iff_right (Nat.le_of_succ_le ha)] + rw [Nat.lt_div_iff_mul_lt' h₁, mul_one, tsub_lt_tsub_iff_right (Nat.le_of_succ_le ha)] exact lt_self_pow₀ (Nat.lt_of_succ_le ha) hb private theorem b_id_helper {a b : ℕ} (ha : 2 ≤ a) (hb : 2 < b) : 2 ≤ (a ^ b + 1) / (a + 1) := by From c4263a50a3891aab0e5722825ac4b9733cd417d6 Mon Sep 17 00:00:00 2001 From: Artie Khovanov Date: Mon, 30 Dec 2024 10:07:52 +0000 Subject: [PATCH 035/189] chore: Rename injectivity theorems for consistency (#20329) Moves: Subgroup.toSubmonoid_eq -> Subgroup.toSubmonoid_inj AddSubgroup.toAddSubmonoid_eq -> AddSubgroup.toAddSubmonoid_inj Submodule.toAddSubmonoid_eq -> Submodule.toAddSubmonoid_inj Submodule.toAddSubgroup_eq -> Submodule.toAddSubgroup_inj Submodule.toSubMulAction_eq -> Submodule.toSubMulAction_inj Cardinal.toNat_eq_iff_eq_of_lt_aleph0 -> Cardinal.toNat_inj_of_lt_aleph0 IntermediateField.toSubalgebra_eq_iff -> IntermediateField.toSubalgebra_inj NormedSpace.Dual.toWeakDual_eq_iff -> NormedSpace.Dual.toWeakDual_inj WeakDual.toNormedDual_eq_iff -> WeakDual.toNormedDual_inj RatFunc.toFractionRing_eq_iff -> RatFunc.toFractionRing_inj Cardinal.toPartENat_eq_iff_of_le_aleph0 -> Cardinal.toPartENat_inj_of_le_aleph0 ENat.toENNReal_coe_eq_iff -> ENat.toENNReal_coe_inj Ordinal.toGame_eq_iff -> Ordinal.toGame_inj Ordinal.toPGame_eq_iff -> Ordinal.toPGame_inj DirichletCharacter.toUnitHom_eq_iff -> DirichletCharacter.toUnitHom_inj LieSubalgebra.coe_to_submodule_eq_iff -> LieSubalgebra.coe_to_submodule_inj LieSubmodule.coe_toSubmodule_eq_iff -> LieSubmodule.coe_toSubmodule_inj --- Mathlib/Algebra/Group/Subgroup/Basic.lean | 2 +- Mathlib/Algebra/Group/Subgroup/Defs.lean | 4 +- Mathlib/Algebra/Group/Subgroup/Pointwise.lean | 2 +- Mathlib/Algebra/Lie/Abelian.lean | 2 +- Mathlib/Algebra/Lie/BaseChange.lean | 2 +- Mathlib/Algebra/Lie/CartanSubalgebra.lean | 4 +- Mathlib/Algebra/Lie/Derivation/Killing.lean | 2 +- Mathlib/Algebra/Lie/Engel.lean | 6 +-- Mathlib/Algebra/Lie/IdealOperations.lean | 6 +-- Mathlib/Algebra/Lie/Nilpotent.lean | 6 +-- Mathlib/Algebra/Lie/Subalgebra.lean | 8 ++-- Mathlib/Algebra/Lie/Submodule.lean | 48 ++++++++++--------- Mathlib/Algebra/Lie/TensorProduct.lean | 2 +- Mathlib/Algebra/Lie/Weights/Basic.lean | 6 +-- Mathlib/Algebra/Lie/Weights/Cartan.lean | 2 +- Mathlib/Algebra/Lie/Weights/Killing.lean | 4 +- Mathlib/Algebra/Lie/Weights/Linear.lean | 4 +- Mathlib/Algebra/Module/Submodule/Defs.lean | 14 ++++-- Mathlib/Algebra/Module/Submodule/Lattice.lean | 2 +- .../Module/Submodule/RestrictScalars.lean | 2 +- Mathlib/Analysis/Normed/Module/WeakDual.lean | 12 +++-- Mathlib/Data/Real/ENatENNReal.lean | 4 +- Mathlib/Deprecated/Cardinal/PartENat.lean | 5 +- .../TopologicalEntropy/CoverEntropy.lean | 2 +- Mathlib/FieldTheory/Adjoin.lean | 4 +- .../IntermediateField/Algebraic.lean | 5 -- .../FieldTheory/IntermediateField/Basic.lean | 12 ++++- Mathlib/FieldTheory/RatFunc/Defs.lean | 4 +- .../GroupTheory/GroupAction/SubMulAction.lean | 2 +- .../DirichletCharacter/Basic.lean | 4 +- Mathlib/SetTheory/Cardinal/ToNat.lean | 4 +- Mathlib/SetTheory/Game/Ordinal.lean | 8 +++- 32 files changed, 115 insertions(+), 79 deletions(-) diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index cb640dadca280..e3e7ad321eee8 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -138,7 +138,7 @@ theorem prod_le_iff {H : Subgroup G} {K : Subgroup N} {J : Subgroup (G × N)} : @[to_additive (attr := simp) prod_eq_bot_iff] theorem prod_eq_bot_iff {H : Subgroup G} {K : Subgroup N} : H.prod K = ⊥ ↔ H = ⊥ ∧ K = ⊥ := by - simpa only [← Subgroup.toSubmonoid_eq] using Submonoid.prod_eq_bot_iff + simpa only [← Subgroup.toSubmonoid_inj] using Submonoid.prod_eq_bot_iff @[to_additive closure_prod] theorem closure_prod {s : Set G} {t : Set N} (hs : 1 ∈ s) (ht : 1 ∈ t) : diff --git a/Mathlib/Algebra/Group/Subgroup/Defs.lean b/Mathlib/Algebra/Group/Subgroup/Defs.lean index 35cda31332cc7..9b08ce6457c62 100644 --- a/Mathlib/Algebra/Group/Subgroup/Defs.lean +++ b/Mathlib/Algebra/Group/Subgroup/Defs.lean @@ -330,9 +330,11 @@ theorem toSubmonoid_injective : Function.Injective (toSubmonoid : Subgroup G → exact SetLike.ext'_iff.2 this @[to_additive (attr := simp)] -theorem toSubmonoid_eq {p q : Subgroup G} : p.toSubmonoid = q.toSubmonoid ↔ p = q := +theorem toSubmonoid_inj {p q : Subgroup G} : p.toSubmonoid = q.toSubmonoid ↔ p = q := toSubmonoid_injective.eq_iff +@[to_additive, deprecated (since := "2024-12-29")] alias toSubmonoid_eq := toSubmonoid_inj + @[to_additive (attr := mono)] theorem toSubmonoid_strictMono : StrictMono (toSubmonoid : Subgroup G → Submonoid G) := fun _ _ => id diff --git a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean index 59e702a1824e6..9d32674ef1a3b 100644 --- a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean +++ b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean @@ -142,7 +142,7 @@ theorem closure_induction_right {p : (x : G) → x ∈ closure s → Prop} (one @[to_additive (attr := simp)] theorem closure_inv (s : Set G) : closure s⁻¹ = closure s := by - simp only [← toSubmonoid_eq, closure_toSubmonoid, inv_inv, union_comm] + simp only [← toSubmonoid_inj, closure_toSubmonoid, inv_inv, union_comm] @[to_additive (attr := simp)] lemma closure_singleton_inv (x : G) : closure {x⁻¹} = closure {x} := by diff --git a/Mathlib/Algebra/Lie/Abelian.lean b/Mathlib/Algebra/Lie/Abelian.lean index 1c88f3135e544..3f1dcaf33206a 100644 --- a/Mathlib/Algebra/Lie/Abelian.lean +++ b/Mathlib/Algebra/Lie/Abelian.lean @@ -121,7 +121,7 @@ instance : IsTrivial L (maxTrivSubmodule R L M) where trivial x m := Subtype.ext @[simp] theorem ideal_oper_maxTrivSubmodule_eq_bot (I : LieIdeal R L) : ⁅I, maxTrivSubmodule R L M⁆ = ⊥ := by - rw [← LieSubmodule.coe_toSubmodule_eq_iff, LieSubmodule.lieIdeal_oper_eq_linear_span, + rw [← LieSubmodule.coe_toSubmodule_inj, LieSubmodule.lieIdeal_oper_eq_linear_span, LieSubmodule.bot_coeSubmodule, Submodule.span_eq_bot] rintro m ⟨⟨x, hx⟩, ⟨⟨m, hm⟩, rfl⟩⟩ exact hm x diff --git a/Mathlib/Algebra/Lie/BaseChange.lean b/Mathlib/Algebra/Lie/BaseChange.lean index 5ca3c13673a21..5a8ada230004c 100644 --- a/Mathlib/Algebra/Lie/BaseChange.lean +++ b/Mathlib/Algebra/Lie/BaseChange.lean @@ -214,7 +214,7 @@ lemma lie_baseChange {I : LieIdeal R L} {N : LieSubmodule R L M} : ⁅I, N⁆.baseChange A = ⁅I.baseChange A, N.baseChange A⁆ := by set s : Set (A ⊗[R] M) := { m | ∃ x ∈ I, ∃ n ∈ N, 1 ⊗ₜ ⁅x, n⁆ = m} have : (TensorProduct.mk R A M 1) '' {m | ∃ x ∈ I, ∃ n ∈ N, ⁅x, n⁆ = m} = s := by ext; simp [s] - rw [← coe_toSubmodule_eq_iff, coe_baseChange, lieIdeal_oper_eq_linear_span', + rw [← coe_toSubmodule_inj, coe_baseChange, lieIdeal_oper_eq_linear_span', Submodule.baseChange_span, this, lieIdeal_oper_eq_linear_span'] refine le_antisymm (Submodule.span_mono ?_) (Submodule.span_le.mpr ?_) · rintro - ⟨x, hx, m, hm, rfl⟩ diff --git a/Mathlib/Algebra/Lie/CartanSubalgebra.lean b/Mathlib/Algebra/Lie/CartanSubalgebra.lean index c904d1921c59f..fed2b9e12669c 100644 --- a/Mathlib/Algebra/Lie/CartanSubalgebra.lean +++ b/Mathlib/Algebra/Lie/CartanSubalgebra.lean @@ -53,7 +53,7 @@ instance [H.IsCartanSubalgebra] : LieAlgebra.IsNilpotent R H := @[simp] theorem normalizer_eq_self_of_isCartanSubalgebra (H : LieSubalgebra R L) [H.IsCartanSubalgebra] : H.toLieSubmodule.normalizer = H.toLieSubmodule := by - rw [← LieSubmodule.coe_toSubmodule_eq_iff, coe_normalizer_eq_normalizer, + rw [← LieSubmodule.coe_toSubmodule_inj, coe_normalizer_eq_normalizer, IsCartanSubalgebra.self_normalizing, coe_toLieSubmodule] @[simp] @@ -84,7 +84,7 @@ theorem isCartanSubalgebra_iff_isUcsLimit : H.IsCartanSubalgebra ↔ H.toLieSubm self_normalizing := by have hk' := hk (k + 1) k.le_succ rw [LieSubmodule.ucs_succ, hk k (le_refl k)] at hk' - rw [← LieSubalgebra.coe_to_submodule_eq_iff, ← LieSubalgebra.coe_normalizer_eq_normalizer, + rw [← LieSubalgebra.coe_to_submodule_inj, ← LieSubalgebra.coe_normalizer_eq_normalizer, hk', LieSubalgebra.coe_toLieSubmodule] } lemma ne_bot_of_isCartanSubalgebra [Nontrivial L] (H : LieSubalgebra R L) [H.IsCartanSubalgebra] : diff --git a/Mathlib/Algebra/Lie/Derivation/Killing.lean b/Mathlib/Algebra/Lie/Derivation/Killing.lean index 65e346eb954d5..04b2110b0fe93 100644 --- a/Mathlib/Algebra/Lie/Derivation/Killing.lean +++ b/Mathlib/Algebra/Lie/Derivation/Killing.lean @@ -94,7 +94,7 @@ lemma killingForm_restrict_range_ad_nondegenerate : /-- The range of the adjoint action on a finite-dimensional Killing Lie algebra is full. -/ @[simp] lemma range_ad_eq_top : 𝕀 = ⊤ := by - rw [← LieSubalgebra.coe_to_submodule_eq_iff] + rw [← LieSubalgebra.coe_to_submodule_inj] apply LinearMap.BilinForm.eq_top_of_restrict_nondegenerate_of_orthogonal_eq_bot (LieModule.traceForm_isSymm R 𝔻 𝔻).isRefl (killingForm_restrict_range_ad_nondegenerate R L) refine (Submodule.eq_bot_iff _).mpr fun D hD ↦ ext fun x ↦ ?_ diff --git a/Mathlib/Algebra/Lie/Engel.lean b/Mathlib/Algebra/Lie/Engel.lean index 8cb392186938b..5bf6908043e53 100644 --- a/Mathlib/Algebra/Lie/Engel.lean +++ b/Mathlib/Algebra/Lie/Engel.lean @@ -128,7 +128,7 @@ theorem isNilpotentOfIsNilpotentSpanSupEqTop (hnp : IsNilpotent <| toEnd R L M x obtain ⟨n, hn⟩ := hnp obtain ⟨k, hk⟩ := hIM have hk' : I.lcs M k = ⊥ := by - simp only [← coe_toSubmodule_eq_iff, I.coe_lcs_eq, hk, bot_coeSubmodule] + simp only [← coe_toSubmodule_inj, I.coe_lcs_eq, hk, bot_coeSubmodule] suffices ∀ l, lowerCentralSeries R L M (l * n) ≤ I.lcs M l by use k * n simpa [hk'] using this k @@ -238,9 +238,9 @@ theorem LieAlgebra.isEngelian_of_isNoetherian [IsNoetherian R L] : LieAlgebra.Is LieSubmodule.nontrivial_iff_ne_bot R K] have : Nontrivial (L' ⧸ K.toLieSubmodule) := by replace hK₂ : K.toLieSubmodule ≠ ⊤ := by - rwa [Ne, ← LieSubmodule.coe_toSubmodule_eq_iff, K.coe_toLieSubmodule, + rwa [Ne, ← LieSubmodule.coe_toSubmodule_inj, K.coe_toLieSubmodule, LieSubmodule.top_coeSubmodule, ← LieSubalgebra.top_coe_submodule, - K.coe_to_submodule_eq_iff] + K.coe_to_submodule_inj] exact Submodule.Quotient.nontrivial_of_lt_top _ hK₂.lt_top have : LieModule.IsNilpotent R K (L' ⧸ K.toLieSubmodule) := by -- Porting note: was refine' hK₁ _ fun x => _ diff --git a/Mathlib/Algebra/Lie/IdealOperations.lean b/Mathlib/Algebra/Lie/IdealOperations.lean index 459dce3dfdbf0..4fd1e86e29fbf 100644 --- a/Mathlib/Algebra/Lie/IdealOperations.lean +++ b/Mathlib/Algebra/Lie/IdealOperations.lean @@ -61,7 +61,7 @@ theorem comap_map_eq (hf : f.ker = ⊥) : comap f (map f N) = N := by @[simp] theorem map_comap_incl : map N.incl (comap N.incl N') = N ⊓ N' := by - rw [← coe_toSubmodule_eq_iff] + rw [← coe_toSubmodule_inj] exact (N : Submodule R M).map_comap_subtype N' variable [LieAlgebra R L] [LieModule R L M₂] (I J : LieIdeal R L) @@ -201,7 +201,7 @@ theorem inf_lie : ⁅I ⊓ J, N⁆ ≤ ⁅I, N⁆ ⊓ ⁅J, N⁆ := by apply mono_lie_left <;> [exact inf_le_left; exact inf_le_right] 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, + rw [← coe_toSubmodule_inj, coeSubmodule_map, lieIdeal_oper_eq_linear_span, lieIdeal_oper_eq_linear_span, Submodule.map_span] congr ext m @@ -266,7 +266,7 @@ theorem map_comap_incl {I₁ I₂ : LieIdeal R L} : map I₁.incl (comap I₁.in theorem comap_bracket_eq {J₁ J₂ : LieIdeal R L'} (h : f.IsIdealMorphism) : comap f ⁅f.idealRange ⊓ J₁, f.idealRange ⊓ J₂⁆ = ⁅comap f J₁, comap f J₂⁆ ⊔ f.ker := by - rw [← LieSubmodule.coe_toSubmodule_eq_iff, comap_coeSubmodule, + rw [← LieSubmodule.coe_toSubmodule_inj, comap_coeSubmodule, LieSubmodule.sup_coe_toSubmodule, f.ker_coeSubmodule, ← Submodule.comap_map_eq, LieSubmodule.lieIdeal_oper_eq_linear_span, LieSubmodule.lieIdeal_oper_eq_linear_span, LinearMap.map_span] diff --git a/Mathlib/Algebra/Lie/Nilpotent.lean b/Mathlib/Algebra/Lie/Nilpotent.lean index d5cad6f1b5f47..96c142a2e09b1 100644 --- a/Mathlib/Algebra/Lie/Nilpotent.lean +++ b/Mathlib/Algebra/Lie/Nilpotent.lean @@ -442,7 +442,7 @@ theorem coe_lcs_range_toEnd_eq (k : ℕ) : theorem isNilpotent_range_toEnd_iff : IsNilpotent R (toEnd R L M).range M ↔ IsNilpotent R L M := by constructor <;> rintro ⟨k, hk⟩ <;> use k <;> - rw [← LieSubmodule.coe_toSubmodule_eq_iff] at hk ⊢ <;> + rw [← LieSubmodule.coe_toSubmodule_inj] at hk ⊢ <;> simpa using hk end LieModule @@ -570,7 +570,7 @@ 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 - rw [← LieSubmodule.coe_toSubmodule_eq_iff] at hk ⊢ + rw [← LieSubmodule.coe_toSubmodule_inj] at hk ⊢ simp [← hf.lieModule_lcs_map_eq hg hfg k, hk] theorem Equiv.lieModule_isNilpotent_iff (f : L ≃ₗ⁅R⁆ L₂) (g : M ≃ₗ[R] M₂) @@ -671,7 +671,7 @@ theorem LieAlgebra.nilpotent_of_nilpotent_quotient {I : LieIdeal R L} (h₁ : I exact LieModule.nilpotentOfNilpotentQuotient R L L h₁ this obtain ⟨k, hk⟩ := h₂ use k - simp [← LieSubmodule.coe_toSubmodule_eq_iff, coe_lowerCentralSeries_ideal_quot_eq, hk] + simp [← LieSubmodule.coe_toSubmodule_inj, coe_lowerCentralSeries_ideal_quot_eq, hk] theorem LieAlgebra.non_trivial_center_of_isNilpotent [Nontrivial L] [IsNilpotent R L] : Nontrivial <| center R L := diff --git a/Mathlib/Algebra/Lie/Subalgebra.lean b/Mathlib/Algebra/Lie/Subalgebra.lean index 7b5748a09788e..4c596fc5e43c6 100644 --- a/Mathlib/Algebra/Lie/Subalgebra.lean +++ b/Mathlib/Algebra/Lie/Subalgebra.lean @@ -194,10 +194,12 @@ theorem to_submodule_injective : Function.Injective ((↑) : LieSubalgebra R L exact h @[simp] -theorem coe_to_submodule_eq_iff (L₁' L₂' : LieSubalgebra R L) : +theorem coe_to_submodule_inj (L₁' L₂' : LieSubalgebra R L) : (L₁' : Submodule R L) = (L₂' : Submodule R L) ↔ L₁' = L₂' := to_submodule_injective.eq_iff +@[deprecated (since := "2024-12-29")] alias coe_to_submodule_eq_iff := coe_to_submodule_inj + theorem coe_to_submodule : ((L' : Submodule R L) : Set L) = L' := rfl @@ -327,7 +329,7 @@ variable (K K' : LieSubalgebra R L) (K₂ : LieSubalgebra R L₂) @[simp] theorem incl_range : K.incl.range = K := by - rw [← coe_to_submodule_eq_iff] + rw [← coe_to_submodule_inj] exact (K : Submodule R L).range_subtype /-- The image of a Lie subalgebra under a Lie algebra morphism is a Lie subalgebra of the @@ -623,7 +625,7 @@ theorem coe_lieSpan_submodule_eq_iff {p : Submodule R L} : · intro x m hm rw [← h, mem_coe_submodule] exact lie_mem _ (subset_lieSpan hm) - · rw [← coe_to_submodule_mk p @h, coe_to_submodule, coe_to_submodule_eq_iff, lieSpan_eq] + · rw [← coe_to_submodule_mk p @h, coe_to_submodule, coe_to_submodule_inj, lieSpan_eq] variable (R L) diff --git a/Mathlib/Algebra/Lie/Submodule.lean b/Mathlib/Algebra/Lie/Submodule.lean index 3408e81f11011..5c39a968185a7 100644 --- a/Mathlib/Algebra/Lie/Submodule.lean +++ b/Mathlib/Algebra/Lie/Submodule.lean @@ -130,9 +130,11 @@ theorem ext (h : ∀ m, m ∈ N ↔ m ∈ N') : N = N' := SetLike.ext h @[simp] -theorem coe_toSubmodule_eq_iff : (N : Submodule R M) = (N' : Submodule R M) ↔ N = N' := +theorem coe_toSubmodule_inj : (N : Submodule R M) = (N' : Submodule R M) ↔ N = N' := coeSubmodule_injective.eq_iff +@[deprecated (since := "2024-12-29")] alias coe_toSubmodule_eq_iff := coe_toSubmodule_inj + /-- Copy of a `LieSubmodule` with a new `carrier` equal to the old one. Useful to fix definitional equalities. -/ protected def copy (s : Set M) (hs : s = ↑N) : LieSubmodule R L M where @@ -198,7 +200,7 @@ instance instLieModule : LieModule R L N where 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 _ _)⟩ + ⟨⟨0⟩, fun _ ↦ (coe_toSubmodule_inj _ _).mp (Subsingleton.elim _ _)⟩ end LieSubmodule @@ -286,7 +288,7 @@ theorem mem_toLieSubmodule (x : L) : x ∈ K.toLieSubmodule ↔ x ∈ K := theorem exists_lieIdeal_coe_eq_iff : (∃ I : LieIdeal R L, ↑I = K) ↔ ∀ x y : L, y ∈ K → ⁅x, y⁆ ∈ K := by - simp only [← coe_to_submodule_eq_iff, LieIdeal.coe_to_lieSubalgebra_to_submodule, + simp only [← coe_to_submodule_inj, LieIdeal.coe_to_lieSubalgebra_to_submodule, Submodule.exists_lieSubmodule_coe_eq_iff L] exact Iff.rfl @@ -335,11 +337,11 @@ theorem bot_coeSubmodule : ((⊥ : LieSubmodule R L M) : Submodule R M) = ⊥ := @[simp] theorem coeSubmodule_eq_bot_iff : (N : Submodule R M) = ⊥ ↔ N = ⊥ := by - rw [← coe_toSubmodule_eq_iff, bot_coeSubmodule] + rw [← coe_toSubmodule_inj, bot_coeSubmodule] @[simp] theorem mk_eq_bot_iff {N : Submodule R M} {h} : (⟨N, h⟩ : LieSubmodule R L M) = ⊥ ↔ N = ⊥ := by - rw [← coe_toSubmodule_eq_iff, bot_coeSubmodule] + rw [← coe_toSubmodule_inj, bot_coeSubmodule] @[simp] theorem mem_bot (x : M) : x ∈ (⊥ : LieSubmodule R L M) ↔ x = 0 := @@ -358,11 +360,11 @@ theorem top_coeSubmodule : ((⊤ : LieSubmodule R L M) : Submodule R M) = ⊤ := @[simp] theorem coeSubmodule_eq_top_iff : (N : Submodule R M) = ⊤ ↔ N = ⊤ := by - rw [← coe_toSubmodule_eq_iff, top_coeSubmodule] + rw [← coe_toSubmodule_inj, top_coeSubmodule] @[simp] theorem mk_eq_top_iff {N : Submodule R M} {h} : (⟨N, h⟩ : LieSubmodule R L M) = ⊤ ↔ N = ⊤ := by - rw [← coe_toSubmodule_eq_iff, top_coeSubmodule] + rw [← coe_toSubmodule_inj, top_coeSubmodule] @[simp] theorem mem_top (x : M) : x ∈ (⊤ : LieSubmodule R L M) := @@ -502,12 +504,12 @@ theorem iSup_induction' {ι} (N : ι → LieSubmodule R L M) {C : (x : M) → (x theorem disjoint_iff_coe_toSubmodule : Disjoint N N' ↔ Disjoint (N : Submodule R M) (N' : Submodule R M) := by - rw [disjoint_iff, disjoint_iff, ← coe_toSubmodule_eq_iff, inf_coe_toSubmodule, bot_coeSubmodule, + rw [disjoint_iff, disjoint_iff, ← coe_toSubmodule_inj, inf_coe_toSubmodule, bot_coeSubmodule, ← disjoint_iff] theorem codisjoint_iff_coe_toSubmodule : Codisjoint N N' ↔ Codisjoint (N : Submodule R M) (N' : Submodule R M) := by - rw [codisjoint_iff, codisjoint_iff, ← coe_toSubmodule_eq_iff, sup_coe_toSubmodule, + rw [codisjoint_iff, codisjoint_iff, ← coe_toSubmodule_inj, sup_coe_toSubmodule, top_coeSubmodule, ← codisjoint_iff] theorem isCompl_iff_coe_toSubmodule : @@ -523,7 +525,7 @@ alias independent_iff_coe_toSubmodule := iSupIndep_iff_coe_toSubmodule theorem iSup_eq_top_iff_coe_toSubmodule {ι : Sort*} {N : ι → LieSubmodule R L M} : ⨆ i, N i = ⊤ ↔ ⨆ i, (N i : Submodule R M) = ⊤ := by - rw [← iSup_coe_toSubmodule, ← top_coeSubmodule (L := L), coe_toSubmodule_eq_iff] + rw [← iSup_coe_toSubmodule, ← top_coeSubmodule (L := L), coe_toSubmodule_inj] instance : Add (LieSubmodule R L M) where add := max @@ -581,7 +583,7 @@ instance [IsArtinian R M] : IsAtomic (LieSubmodule R L M) := @[simp] theorem subsingleton_iff : Subsingleton (LieSubmodule R L M) ↔ Subsingleton M := have h : Subsingleton (LieSubmodule R L M) ↔ Subsingleton (Submodule R M) := by - rw [← subsingleton_iff_bot_eq_top, ← subsingleton_iff_bot_eq_top, ← coe_toSubmodule_eq_iff, + rw [← subsingleton_iff_bot_eq_top, ← subsingleton_iff_bot_eq_top, ← coe_toSubmodule_inj, top_coeSubmodule, bot_coeSubmodule] h.trans <| Submodule.subsingleton_iff R @@ -685,7 +687,7 @@ theorem coe_lieSpan_submodule_eq_iff {p : Submodule R M} : (lieSpan R L (p : Set M) : Submodule R M) = p ↔ ∃ N : LieSubmodule R L M, ↑N = p := by rw [p.exists_lieSubmodule_coe_eq_iff L]; constructor <;> intro h · intro x m hm; rw [← h, mem_coeSubmodule]; exact lie_mem _ (subset_lieSpan hm) - · rw [← coe_toSubmodule_mk p @h, coe_toSubmodule, coe_toSubmodule_eq_iff, lieSpan_eq] + · rw [← coe_toSubmodule_mk p @h, coe_toSubmodule, coe_toSubmodule_inj, lieSpan_eq] variable (R L M) @@ -830,11 +832,11 @@ theorem mem_comap {m : M} : m ∈ comap f N' ↔ f m ∈ N' := Iff.rfl theorem comap_incl_eq_top : N₂.comap N.incl = ⊤ ↔ N ≤ N₂ := by - rw [← LieSubmodule.coe_toSubmodule_eq_iff, LieSubmodule.coeSubmodule_comap, LieSubmodule.incl_coe, + rw [← LieSubmodule.coe_toSubmodule_inj, LieSubmodule.coeSubmodule_comap, LieSubmodule.incl_coe, LieSubmodule.top_coeSubmodule, Submodule.comap_subtype_eq_top, coeSubmodule_le_coeSubmodule] theorem comap_incl_eq_bot : N₂.comap N.incl = ⊥ ↔ N ⊓ N₂ = ⊥ := by - simp only [← coe_toSubmodule_eq_iff, coeSubmodule_comap, incl_coe, bot_coeSubmodule, + simp only [← coe_toSubmodule_inj, coeSubmodule_comap, incl_coe, bot_coeSubmodule, inf_coe_toSubmodule] rw [← Submodule.disjoint_iff_comap_eq_bot, disjoint_iff] @@ -1030,7 +1032,7 @@ theorem IsIdealMorphism.eq (hf : f.IsIdealMorphism) : f.idealRange = f.range := theorem isIdealMorphism_iff : f.IsIdealMorphism ↔ ∀ (x : L') (y : L), ∃ z : L, ⁅x, f y⁆ = f z := by simp only [isIdealMorphism_def, idealRange_eq_lieSpan_range, ← - LieSubalgebra.coe_to_submodule_eq_iff, ← f.range.coe_to_submodule, + LieSubalgebra.coe_to_submodule_inj, ← f.range.coe_to_submodule, LieIdeal.coe_to_lieSubalgebra_to_submodule, LieSubmodule.coe_lieSpan_submodule_eq_iff, LieSubalgebra.mem_coe_submodule, mem_range, exists_imp, Submodule.exists_lieSubmodule_coe_eq_iff] @@ -1074,7 +1076,7 @@ theorem le_ker_iff : I ≤ f.ker ↔ ∀ x, x ∈ I → f x = 0 := by · rw [mem_ker]; apply h x hx theorem ker_eq_bot : f.ker = ⊥ ↔ Function.Injective f := by - rw [← LieSubmodule.coe_toSubmodule_eq_iff, ker_coeSubmodule, LieSubmodule.bot_coeSubmodule, + rw [← LieSubmodule.coe_toSubmodule_inj, ker_coeSubmodule, LieSubmodule.bot_coeSubmodule, LinearMap.ker_eq_bot, coe_toLinearMap] @[simp] @@ -1082,14 +1084,14 @@ theorem range_coeSubmodule : (f.range : Submodule R L') = LinearMap.range (f : L rfl theorem range_eq_top : f.range = ⊤ ↔ Function.Surjective f := by - rw [← LieSubalgebra.coe_to_submodule_eq_iff, range_coeSubmodule, LieSubalgebra.top_coe_submodule] + rw [← LieSubalgebra.coe_to_submodule_inj, range_coeSubmodule, LieSubalgebra.top_coe_submodule] exact LinearMap.range_eq_top @[simp] theorem idealRange_eq_top_of_surjective (h : Function.Surjective f) : f.idealRange = ⊤ := by rw [← f.range_eq_top] at h rw [idealRange_eq_lieSpan_range, h, ← LieSubalgebra.coe_to_submodule, ← - LieSubmodule.coe_toSubmodule_eq_iff, LieSubmodule.top_coeSubmodule, + LieSubmodule.coe_toSubmodule_inj, LieSubmodule.top_coeSubmodule, LieSubalgebra.top_coe_submodule, LieSubmodule.coe_lieSpan_submodule_eq_iff] use ⊤ exact LieSubmodule.top_coeSubmodule @@ -1182,7 +1184,7 @@ theorem map_comap_eq (h : f.IsIdealMorphism) : map f (comap f J) = f.idealRange @[simp] theorem comap_map_eq (h : ↑(map f I) = f '' I) : comap f (map f I) = I ⊔ f.ker := by - rw [← LieSubmodule.coe_toSubmodule_eq_iff, comap_coeSubmodule, I.map_coeSubmodule f h, + rw [← LieSubmodule.coe_toSubmodule_inj, comap_coeSubmodule, I.map_coeSubmodule f h, LieSubmodule.sup_coe_toSubmodule, f.ker_coeSubmodule, Submodule.comap_map_eq] variable (f I J) @@ -1216,7 +1218,7 @@ theorem ker_incl : I.incl.ker = ⊥ := by ext; simp @[simp] theorem incl_idealRange : I.incl.idealRange = I := by rw [LieHom.idealRange_eq_lieSpan_range, ← LieSubalgebra.coe_to_submodule, ← - LieSubmodule.coe_toSubmodule_eq_iff, incl_range, coe_to_lieSubalgebra_to_submodule, + LieSubmodule.coe_toSubmodule_inj, incl_range, coe_to_lieSubalgebra_to_submodule, LieSubmodule.coe_lieSpan_submodule_eq_iff] use I @@ -1245,7 +1247,7 @@ theorem ker_coeSubmodule : (f.ker : Submodule R M) = LinearMap.ker (f : M →ₗ rfl theorem ker_eq_bot : f.ker = ⊥ ↔ Function.Injective f := by - rw [← LieSubmodule.coe_toSubmodule_eq_iff, ker_coeSubmodule, LieSubmodule.bot_coeSubmodule, + rw [← LieSubmodule.coe_toSubmodule_inj, ker_coeSubmodule, LieSubmodule.bot_coeSubmodule, LinearMap.ker_eq_bot, coe_toLinearMap] variable {f} @@ -1316,12 +1318,12 @@ theorem ker_incl : N.incl.ker = ⊥ := (LieModuleHom.ker_eq_bot N.incl).mpr <| i @[simp] theorem range_incl : N.incl.range = N := by - simp only [← coe_toSubmodule_eq_iff, LieModuleHom.coeSubmodule_range, incl_coe] + simp only [← coe_toSubmodule_inj, LieModuleHom.coeSubmodule_range, incl_coe] rw [Submodule.range_subtype] @[simp] theorem comap_incl_self : comap N.incl N = ⊤ := by - simp only [← coe_toSubmodule_eq_iff, coeSubmodule_comap, incl_coe, top_coeSubmodule] + simp only [← coe_toSubmodule_inj, coeSubmodule_comap, incl_coe, top_coeSubmodule] rw [Submodule.comap_subtype_self] theorem map_incl_top : (⊤ : LieSubmodule R L N).map N.incl = N := by simp diff --git a/Mathlib/Algebra/Lie/TensorProduct.lean b/Mathlib/Algebra/Lie/TensorProduct.lean index 0974b6a2d5a2d..759d0c87ded6b 100644 --- a/Mathlib/Algebra/Lie/TensorProduct.lean +++ b/Mathlib/Algebra/Lie/TensorProduct.lean @@ -196,7 +196,7 @@ applying the action of `L` on `M`, we obtain morphism of Lie modules `f : I ⊗ This lemma states that `⁅I, N⁆ = range f`. -/ theorem lieIdeal_oper_eq_tensor_map_range : ⁅I, N⁆ = ((toModuleHom R L M).comp (mapIncl I N : I ⊗[R] N →ₗ⁅R,L⁆ L ⊗[R] M)).range := by - rw [← coe_toSubmodule_eq_iff, lieIdeal_oper_eq_linear_span, LieModuleHom.coeSubmodule_range, + rw [← coe_toSubmodule_inj, lieIdeal_oper_eq_linear_span, LieModuleHom.coeSubmodule_range, LieModuleHom.coe_linearMap_comp, LinearMap.range_comp, mapIncl_def, coe_linearMap_map, TensorProduct.map_range_eq_span_tmul, Submodule.map_span] congr; ext m; constructor diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index f3bb8d4eb7128..4828dcf2ae899 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -589,7 +589,7 @@ variable [IsNoetherian R M] [IsArtinian R M] 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, + simpa only [isCompl_iff, codisjoint_iff, disjoint_iff, ← LieSubmodule.coe_toSubmodule_inj, LieSubmodule.sup_coe_toSubmodule, LieSubmodule.inf_coe_toSubmodule, LieSubmodule.top_coeSubmodule, LieSubmodule.bot_coeSubmodule, coe_genWeightSpaceOf_zero] using (toEnd R L M x).isCompl_iSup_ker_pow_iInf_range_pow @@ -728,7 +728,7 @@ instance [IsTriangularizable R L M] : IsTriangularizable R (LieModule.toEnd R L @[simp] 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, + rw [← LieSubmodule.coe_toSubmodule_inj, LieSubmodule.iSup_coe_toSubmodule, LieSubmodule.top_coeSubmodule] dsimp [genWeightSpaceOf] exact IsTriangularizable.maxGenEigenspace_eq_top x @@ -766,7 +766,7 @@ instance (N : LieSubmodule K L M) [IsTriangularizable K L M] : IsTriangularizabl See also `LieModule.iSup_genWeightSpace_eq_top'`. -/ lemma iSup_genWeightSpace_eq_top [IsTriangularizable K L M] : ⨆ χ : L → K, genWeightSpace M χ = ⊤ := by - simp only [← LieSubmodule.coe_toSubmodule_eq_iff, LieSubmodule.iSup_coe_toSubmodule, + simp only [← LieSubmodule.coe_toSubmodule_inj, LieSubmodule.iSup_coe_toSubmodule, LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.top_coeSubmodule, genWeightSpace] refine Module.End.iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo (toEnd K L M) (fun x y φ z ↦ (genWeightSpaceOf M φ y).lie_mem) ?_ diff --git a/Mathlib/Algebra/Lie/Weights/Cartan.lean b/Mathlib/Algebra/Lie/Weights/Cartan.lean index e494e0e19ba5d..576a1918640c9 100644 --- a/Mathlib/Algebra/Lie/Weights/Cartan.lean +++ b/Mathlib/Algebra/Lie/Weights/Cartan.lean @@ -240,7 +240,7 @@ theorem zeroRootSubalgebra_eq_iff_is_cartan [IsNoetherian R L] : @[simp] theorem rootSpace_zero_eq (H : LieSubalgebra R L) [H.IsCartanSubalgebra] [IsNoetherian R L] : rootSpace H 0 = H.toLieSubmodule := by - rw [← LieSubmodule.coe_toSubmodule_eq_iff, ← coe_zeroRootSubalgebra, + rw [← LieSubmodule.coe_toSubmodule_inj, ← coe_zeroRootSubalgebra, zeroRootSubalgebra_eq_of_is_cartan R L H, LieSubalgebra.coe_toLieSubmodule] variable {R L H} diff --git a/Mathlib/Algebra/Lie/Weights/Killing.lean b/Mathlib/Algebra/Lie/Weights/Killing.lean index a4667d0fba6aa..fa7a40102df4a 100644 --- a/Mathlib/Algebra/Lie/Weights/Killing.lean +++ b/Mathlib/Algebra/Lie/Weights/Killing.lean @@ -54,7 +54,7 @@ lemma ker_restrict_eq_bot_of_isCartanSubalgebra have h : Codisjoint (rootSpace H 0) (LieModule.posFittingComp R H L) := (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, + rwa [codisjoint_iff, ← LieSubmodule.coe_toSubmodule_inj, LieSubmodule.sup_coe_toSubmodule, LieSubmodule.top_coeSubmodule, rootSpace_zero_eq R L H, LieSubalgebra.coe_toLieSubmodule, ← codisjoint_iff] at h suffices this : ∀ m₀ ∈ H, ∀ m₁ ∈ LieModule.posFittingComp R H L, killingForm R L m₀ m₁ = 0 by @@ -424,7 +424,7 @@ lemma isCompl_ker_weight_span_coroot (α : Weight K H L) : apply Module.Dual.isCompl_ker_of_disjoint_of_ne_bot (by aesop) (disjoint_ker_weight_corootSpace α) replace hα : corootSpace α ≠ ⊥ := by simpa using hα - rwa [ne_eq, ← LieSubmodule.coe_toSubmodule_eq_iff] at hα + rwa [ne_eq, ← LieSubmodule.coe_toSubmodule_inj] at hα lemma traceForm_eq_zero_of_mem_ker_of_mem_span_coroot {α : Weight K H L} {x y : H} (hx : x ∈ α.ker) (hy : y ∈ K ∙ coroot α) : diff --git a/Mathlib/Algebra/Lie/Weights/Linear.lean b/Mathlib/Algebra/Lie/Weights/Linear.lean index 5e1934aa2fd39..c2ca4eedbf3b4 100644 --- a/Mathlib/Algebra/Lie/Weights/Linear.lean +++ b/Mathlib/Algebra/Lie/Weights/Linear.lean @@ -97,13 +97,13 @@ instance instLinearWeightsOfIsLieAbelian [IsLieAbelian L] [NoZeroSMulDivisors R 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, genWeightSpace, genWeightSpaceOf, + simp_rw [Ne, ← LieSubmodule.coe_toSubmodule_inj, 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, genWeightSpace, genWeightSpaceOf, + simp_rw [Ne, ← LieSubmodule.coe_toSubmodule_inj, 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 diff --git a/Mathlib/Algebra/Module/Submodule/Defs.lean b/Mathlib/Algebra/Module/Submodule/Defs.lean index 7c2c0a6fdb1eb..572cdc0af7b5b 100644 --- a/Mathlib/Algebra/Module/Submodule/Defs.lean +++ b/Mathlib/Algebra/Module/Submodule/Defs.lean @@ -84,7 +84,7 @@ theorem mk_le_mk {S S' : AddSubmonoid M} (h h') : theorem ext (h : ∀ x, x ∈ p ↔ x ∈ q) : p = q := SetLike.ext h --- Porting note: adding this as the `simp`-normal form of `toSubMulAction_eq` +-- Porting note: adding this as the `simp`-normal form of `toSubMulAction_inj` @[simp] theorem carrier_inj : p.carrier = q.carrier ↔ p = q := (SetLike.coe_injective (A := Submodule R M)).eq_iff @@ -108,9 +108,11 @@ theorem toAddSubmonoid_injective : Injective (toAddSubmonoid : Submodule R M → fun p q h => SetLike.ext'_iff.2 (show (p.toAddSubmonoid : Set M) = q from SetLike.ext'_iff.1 h) @[simp] -theorem toAddSubmonoid_eq : p.toAddSubmonoid = q.toAddSubmonoid ↔ p = q := +theorem toAddSubmonoid_inj : p.toAddSubmonoid = q.toAddSubmonoid ↔ p = q := toAddSubmonoid_injective.eq_iff +@[deprecated (since := "2024-12-29")] alias toAddSubmonoid_eq := toAddSubmonoid_inj + @[simp] theorem coe_toAddSubmonoid (p : Submodule R M) : (p.toAddSubmonoid : Set M) = p := rfl @@ -118,9 +120,11 @@ theorem coe_toAddSubmonoid (p : Submodule R M) : (p.toAddSubmonoid : Set M) = p theorem toSubMulAction_injective : Injective (toSubMulAction : Submodule R M → SubMulAction R M) := fun p q h => SetLike.ext'_iff.2 (show (p.toSubMulAction : Set M) = q from SetLike.ext'_iff.1 h) -theorem toSubMulAction_eq : p.toSubMulAction = q.toSubMulAction ↔ p = q := +theorem toSubMulAction_inj : p.toSubMulAction = q.toSubMulAction ↔ p = q := toSubMulAction_injective.eq_iff +@[deprecated (since := "2024-12-29")] alias toSubMulAction_eq := toSubMulAction_inj + @[simp] theorem coe_toSubMulAction (p : Submodule R M) : (p.toSubMulAction : Set M) = p := rfl @@ -287,9 +291,11 @@ theorem toAddSubgroup_injective : Injective (toAddSubgroup : Submodule R M → A | _, _, h => SetLike.ext (SetLike.ext_iff.1 h : _) @[simp] -theorem toAddSubgroup_eq : p.toAddSubgroup = p'.toAddSubgroup ↔ p = p' := +theorem toAddSubgroup_inj : p.toAddSubgroup = p'.toAddSubgroup ↔ p = p' := toAddSubgroup_injective.eq_iff +@[deprecated (since := "2024-12-29")] alias toAddSubgroup_eq := toAddSubgroup_inj + protected theorem sub_mem : x ∈ p → y ∈ p → x - y ∈ p := sub_mem diff --git a/Mathlib/Algebra/Module/Submodule/Lattice.lean b/Mathlib/Algebra/Module/Submodule/Lattice.lean index 364a656683578..b36e7847a330b 100644 --- a/Mathlib/Algebra/Module/Submodule/Lattice.lean +++ b/Mathlib/Algebra/Module/Submodule/Lattice.lean @@ -314,7 +314,7 @@ variable (R) @[simp] theorem subsingleton_iff : Subsingleton (Submodule R M) ↔ Subsingleton M := have h : Subsingleton (Submodule R M) ↔ Subsingleton (AddSubmonoid M) := by - rw [← subsingleton_iff_bot_eq_top, ← subsingleton_iff_bot_eq_top, ← toAddSubmonoid_eq, + rw [← subsingleton_iff_bot_eq_top, ← subsingleton_iff_bot_eq_top, ← toAddSubmonoid_inj, bot_toAddSubmonoid, top_toAddSubmonoid] h.trans AddSubmonoid.subsingleton_iff diff --git a/Mathlib/Algebra/Module/Submodule/RestrictScalars.lean b/Mathlib/Algebra/Module/Submodule/RestrictScalars.lean index d0ad383a6005f..54f609912e963 100644 --- a/Mathlib/Algebra/Module/Submodule/RestrictScalars.lean +++ b/Mathlib/Algebra/Module/Submodule/RestrictScalars.lean @@ -106,6 +106,6 @@ we can turn an `R`-submodule into an `S`-submodule by forgetting the action of ` def restrictScalarsLatticeHom : CompleteLatticeHom (Submodule R M) (Submodule S M) where toFun := restrictScalars S map_sInf' s := by ext; simp - map_sSup' s := by rw [← toAddSubmonoid_eq, toAddSubmonoid_sSup, ← Set.image_comp]; simp + map_sSup' s := by rw [← toAddSubmonoid_inj, toAddSubmonoid_sSup, ← Set.image_comp]; simp end Submodule diff --git a/Mathlib/Analysis/Normed/Module/WeakDual.lean b/Mathlib/Analysis/Normed/Module/WeakDual.lean index a5798eaf704c2..a607393f81d14 100644 --- a/Mathlib/Analysis/Normed/Module/WeakDual.lean +++ b/Mathlib/Analysis/Normed/Module/WeakDual.lean @@ -119,8 +119,10 @@ theorem coe_toWeakDual (x' : Dual 𝕜 E) : toWeakDual x' = x' := rfl @[simp] -theorem toWeakDual_eq_iff (x' y' : Dual 𝕜 E) : toWeakDual x' = toWeakDual y' ↔ x' = y' := - Function.Injective.eq_iff <| LinearEquiv.injective toWeakDual +theorem toWeakDual_inj (x' y' : Dual 𝕜 E) : toWeakDual x' = toWeakDual y' ↔ x' = y' := + (LinearEquiv.injective toWeakDual).eq_iff + +@[deprecated (since := "2024-12-29")] alias toWeakDual_eq_iff := toWeakDual_inj theorem toWeakDual_continuous : Continuous fun x' : Dual 𝕜 E => toWeakDual x' := WeakBilin.continuous_of_continuous_eval _ fun z => (inclusionInDoubleDual 𝕜 E z).continuous @@ -160,8 +162,10 @@ theorem coe_toNormedDual (x' : WeakDual 𝕜 E) : toNormedDual x' = x' := rfl @[simp] -theorem toNormedDual_eq_iff (x' y' : WeakDual 𝕜 E) : toNormedDual x' = toNormedDual y' ↔ x' = y' := - Function.Injective.eq_iff <| LinearEquiv.injective toNormedDual +theorem toNormedDual_inj (x' y' : WeakDual 𝕜 E) : toNormedDual x' = toNormedDual y' ↔ x' = y' := + (LinearEquiv.injective toNormedDual).eq_iff + +@[deprecated (since := "2024-12-29")] alias toNormedDual_eq_iff := toNormedDual_inj theorem isClosed_closedBall (x' : Dual 𝕜 E) (r : ℝ) : IsClosed (toNormedDual ⁻¹' closedBall x' r) := isClosed_induced_iff'.2 (ContinuousLinearMap.is_weak_closed_closedBall x' r) diff --git a/Mathlib/Data/Real/ENatENNReal.lean b/Mathlib/Data/Real/ENatENNReal.lean index 69b2552e637db..2e04581738900 100644 --- a/Mathlib/Data/Real/ENatENNReal.lean +++ b/Mathlib/Data/Real/ENatENNReal.lean @@ -57,9 +57,11 @@ theorem toENNReal_ofNat (n : ℕ) [n.AtLeastTwo] : rfl @[simp, norm_cast] -theorem toENNReal_coe_eq_iff : (m : ℝ≥0∞) = (n : ℝ≥0∞) ↔ m = n := +theorem toENNReal_inj : (m : ℝ≥0∞) = (n : ℝ≥0∞) ↔ m = n := toENNRealOrderEmbedding.eq_iff_eq +@[deprecated (since := "2024-12-29")] alias toENNReal_coe_eq_iff := toENNReal_inj + @[simp, norm_cast] theorem toENNReal_le : (m : ℝ≥0∞) ≤ n ↔ m ≤ n := toENNRealOrderEmbedding.le_iff_le diff --git a/Mathlib/Deprecated/Cardinal/PartENat.lean b/Mathlib/Deprecated/Cardinal/PartENat.lean index 8387d51f125e4..e2c263a8cfb2e 100644 --- a/Mathlib/Deprecated/Cardinal/PartENat.lean +++ b/Mathlib/Deprecated/Cardinal/PartENat.lean @@ -79,10 +79,13 @@ lemma toPartENat_le_iff_of_lt_aleph0 {c c' : Cardinal} (hc' : c' < ℵ₀) : simp_rw [← partENatOfENat_toENat, toENat_nat, ← toENat_le_nat, ← PartENat.withTopOrderIso.symm.le_iff_le, PartENat.ofENat_le, map_le_map_iff] -lemma toPartENat_eq_iff_of_le_aleph0 {c c' : Cardinal} (hc : c ≤ ℵ₀) (hc' : c' ≤ ℵ₀) : +lemma toPartENat_inj_of_le_aleph0 {c c' : Cardinal} (hc : c ≤ ℵ₀) (hc' : c' ≤ ℵ₀) : toPartENat c = toPartENat c' ↔ c = c' := toPartENat_strictMonoOn.injOn.eq_iff hc hc' +@[deprecated (since := "2024-12-29")] alias toPartENat_eq_iff_of_le_aleph0 := + toPartENat_inj_of_le_aleph0 + theorem toPartENat_mono {c c' : Cardinal} (h : c ≤ c') : toPartENat c ≤ toPartENat c' := OrderHomClass.mono _ h diff --git a/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean b/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean index 39839f3ed92ac..01e2eea5fab82 100644 --- a/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean +++ b/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean @@ -470,7 +470,7 @@ lemma coverEntropyEntourage_le_log_coverMincard_div {T : X → X} {F : Set X} (F coverEntropyEntourage T F (U ○ U) ≤ log (coverMincard T F U n) / n := by -- Deal with the edge cases: `F = ∅` or `F` has no finite cover. rcases eq_or_ne (log (coverMincard T F U n)) ⊥ with logm_bot | logm_nneg - · rw [log_eq_bot_iff, ← ENat.toENNReal_zero, ENat.toENNReal_coe_eq_iff, + · rw [log_eq_bot_iff, ← ENat.toENNReal_zero, ENat.toENNReal_inj, coverMincard_eq_zero_iff T F U n] at logm_bot simp [logm_bot] rcases eq_or_ne (log (coverMincard T F U n)) ⊤ with logm_top | logm_fin diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index 6c85ffedec641..80352130ce0b4 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -991,12 +991,12 @@ variable {K L : IntermediateField F E} @[simp] theorem rank_eq_one_iff : Module.rank F K = 1 ↔ K = ⊥ := by - rw [← toSubalgebra_eq_iff, ← rank_eq_rank_subalgebra, Subalgebra.rank_eq_one_iff, + rw [← toSubalgebra_inj, ← rank_eq_rank_subalgebra, Subalgebra.rank_eq_one_iff, bot_toSubalgebra] @[simp] theorem finrank_eq_one_iff : finrank F K = 1 ↔ K = ⊥ := by - rw [← toSubalgebra_eq_iff, ← finrank_eq_finrank_subalgebra, Subalgebra.finrank_eq_one_iff, + rw [← toSubalgebra_inj, ← finrank_eq_finrank_subalgebra, Subalgebra.finrank_eq_one_iff, bot_toSubalgebra] @[simp] diff --git a/Mathlib/FieldTheory/IntermediateField/Algebraic.lean b/Mathlib/FieldTheory/IntermediateField/Algebraic.lean index db5ed7f190e80..f134a4e0b15cd 100644 --- a/Mathlib/FieldTheory/IntermediateField/Algebraic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Algebraic.lean @@ -58,11 +58,6 @@ theorem finrank_eq_finrank_subalgebra : finrank K F.toSubalgebra = finrank K F : variable {F} {E} -@[simp] -theorem toSubalgebra_eq_iff : F.toSubalgebra = E.toSubalgebra ↔ F = E := by - rw [SetLike.ext_iff, SetLike.ext'_iff, Set.ext_iff] - rfl - /-- If `F ≤ E` are two intermediate fields of `L / K` such that `[E : K] ≤ [F : K]` are finite, then `F = E`. -/ theorem eq_of_le_of_finrank_le [hfin : FiniteDimensional K E] (h_le : F ≤ E) diff --git a/Mathlib/FieldTheory/IntermediateField/Basic.lean b/Mathlib/FieldTheory/IntermediateField/Basic.lean index a02aac491cc51..a80409efa5208 100644 --- a/Mathlib/FieldTheory/IntermediateField/Basic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Basic.lean @@ -565,10 +565,20 @@ theorem toSubfield_injective : Function.Injective (toSubfield : IntermediateFiel ext simp_rw [← mem_toSubfield, h] +variable {F E : IntermediateField K L} + +@[simp] +theorem toSubalgebra_inj : F.toSubalgebra = E.toSubalgebra ↔ F = E := toSubalgebra_injective.eq_iff + +@[deprecated (since := "2024-12-29")] alias toSubalgebra_eq_iff := toSubalgebra_inj + +@[simp] +theorem toSubfield_inj : F.toSubfield = E.toSubfield ↔ F = E := toSubfield_injective.eq_iff + theorem map_injective (f : L →ₐ[K] L') : Function.Injective (map f) := by intro _ _ h rwa [← toSubalgebra_injective.eq_iff, toSubalgebra_map, toSubalgebra_map, - (Subalgebra.map_injective f.injective).eq_iff, toSubalgebra_injective.eq_iff] at h + (Subalgebra.map_injective f.injective).eq_iff, toSubalgebra_inj] at h variable (S) diff --git a/Mathlib/FieldTheory/RatFunc/Defs.lean b/Mathlib/FieldTheory/RatFunc/Defs.lean index 7579e837d384f..a0e7e1e83de5b 100644 --- a/Mathlib/FieldTheory/RatFunc/Defs.lean +++ b/Mathlib/FieldTheory/RatFunc/Defs.lean @@ -83,10 +83,12 @@ theorem ofFractionRing_injective : Function.Injective (ofFractionRing : _ → Ra theorem toFractionRing_injective : Function.Injective (toFractionRing : _ → FractionRing K[X]) | ⟨x⟩, ⟨y⟩, xy => by subst xy; rfl -@[simp] lemma toFractionRing_eq_iff {x y : RatFunc K} : +@[simp] lemma toFractionRing_inj {x y : RatFunc K} : toFractionRing x = toFractionRing y ↔ x = y := toFractionRing_injective.eq_iff +@[deprecated (since := "2024-12-29")] alias toFractionRing_eq_iff := toFractionRing_inj + /-- Non-dependent recursion principle for `RatFunc K`: To construct a term of `P : Sort*` out of `x : RatFunc K`, it suffices to provide a constructor `f : Π (p q : K[X]), P` diff --git a/Mathlib/GroupTheory/GroupAction/SubMulAction.lean b/Mathlib/GroupTheory/GroupAction/SubMulAction.lean index ea413a276ea9c..babd428140d21 100644 --- a/Mathlib/GroupTheory/GroupAction/SubMulAction.lean +++ b/Mathlib/GroupTheory/GroupAction/SubMulAction.lean @@ -351,7 +351,7 @@ lemma orbitRel_of_subMul (p : SubMulAction R M) : /-- Stabilizers in group SubMulAction coincide with stabilizers in the ambient space -/ theorem stabilizer_of_subMul {p : SubMulAction R M} (m : p) : MulAction.stabilizer R m = MulAction.stabilizer R (m : M) := by - rw [← Subgroup.toSubmonoid_eq] + rw [← Subgroup.toSubmonoid_inj] exact stabilizer_of_subMul.submonoid m end MulActionGroup diff --git a/Mathlib/NumberTheory/DirichletCharacter/Basic.lean b/Mathlib/NumberTheory/DirichletCharacter/Basic.lean index 6e8845121d963..2cc4463a984a0 100644 --- a/Mathlib/NumberTheory/DirichletCharacter/Basic.lean +++ b/Mathlib/NumberTheory/DirichletCharacter/Basic.lean @@ -41,7 +41,9 @@ namespace DirichletCharacter lemma toUnitHom_eq_char' {a : ZMod n} (ha : IsUnit a) : χ a = χ.toUnitHom ha.unit := by simp -lemma toUnitHom_eq_iff (ψ : DirichletCharacter R n) : toUnitHom χ = toUnitHom ψ ↔ χ = ψ := by simp +lemma toUnitHom_inj (ψ : DirichletCharacter R n) : toUnitHom χ = toUnitHom ψ ↔ χ = ψ := by simp + +@[deprecated (since := "2024-12-29")] alias toUnitHom_eq_iff := toUnitHom_inj lemma eval_modulus_sub (x : ZMod n) : χ (n - x) = χ (-x) := by simp diff --git a/Mathlib/SetTheory/Cardinal/ToNat.lean b/Mathlib/SetTheory/Cardinal/ToNat.lean index 44953f885edf7..0763034be27cb 100644 --- a/Mathlib/SetTheory/Cardinal/ToNat.lean +++ b/Mathlib/SetTheory/Cardinal/ToNat.lean @@ -64,10 +64,12 @@ theorem toNat_injOn : InjOn toNat (Iio ℵ₀) := toNat_strictMonoOn.injOn /-- Two finite cardinals are equal iff they are equal their `Cardinal.toNat` projections are equal. -/ -theorem toNat_eq_iff_eq_of_lt_aleph0 (hc : c < ℵ₀) (hd : d < ℵ₀) : +theorem toNat_inj_of_lt_aleph0 (hc : c < ℵ₀) (hd : d < ℵ₀) : toNat c = toNat d ↔ c = d := toNat_injOn.eq_iff hc hd +@[deprecated (since := "2024-12-29")] alias toNat_eq_iff_eq_of_lt_aleph0 := toNat_inj_of_lt_aleph0 + theorem toNat_le_iff_le_of_lt_aleph0 (hc : c < ℵ₀) (hd : d < ℵ₀) : toNat c ≤ toNat d ↔ c ≤ d := toNat_strictMonoOn.le_iff_le hc hd diff --git a/Mathlib/SetTheory/Game/Ordinal.lean b/Mathlib/SetTheory/Game/Ordinal.lean index cca43ad7f0086..e770e5e3afd5d 100644 --- a/Mathlib/SetTheory/Game/Ordinal.lean +++ b/Mathlib/SetTheory/Game/Ordinal.lean @@ -144,9 +144,11 @@ theorem toPGame_injective : Function.Injective Ordinal.toPGame := fun _ _ h => toPGame_equiv_iff.1 <| equiv_of_eq h @[simp] -theorem toPGame_eq_iff {a b : Ordinal} : a.toPGame = b.toPGame ↔ a = b := +theorem toPGame_inj {a b : Ordinal} : a.toPGame = b.toPGame ↔ a = b := toPGame_injective.eq_iff +@[deprecated (since := "2024-12-29")] alias toPGame_eq_iff := toPGame_inj + /-- The order embedding version of `toPGame`. -/ @[simps] noncomputable def toPGameEmbedding : Ordinal.{u} ↪o PGame.{u} where @@ -187,9 +189,11 @@ theorem toGame_le_iff {a b : Ordinal} : toGame a ≤ toGame b ↔ a ≤ b := theorem toGame_lt_iff {a b : Ordinal} : toGame a < toGame b ↔ a < b := toPGame_lt_iff -theorem toGame_eq_iff {a b : Ordinal} : toGame a = toGame b ↔ a = b := +theorem toGame_inj {a b : Ordinal} : toGame a = toGame b ↔ a = b := toGameEmbedding.inj +@[deprecated (since := "2024-12-29")] alias toGame_eq_iff := toGame_inj + /-- The natural addition of ordinals corresponds to their sum as games. -/ theorem toPGame_nadd (a b : Ordinal) : (a ♯ b).toPGame ≈ a.toPGame + b.toPGame := by refine ⟨le_of_forall_lf (fun i => ?_) isEmptyElim, le_of_forall_lf (fun i => ?_) isEmptyElim⟩ From 419dc33917cf1df9294863e57154c226ee5bf367 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Mon, 30 Dec 2024 14:39:22 +0000 Subject: [PATCH 036/189] feat(Topology/Algebra): add `TopologicalGroup.isInducing_iff_nhds_one` (#20281) --- Mathlib/Topology/Algebra/Group/Basic.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 6bbe410139376..003e76adb5395 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -811,6 +811,17 @@ theorem continuous_of_continuousAt_one₂ {H M : Type*} [CommMonoid M] [Topologi (((hl x).comp tendsto_snd).mul hf)).mono_right (le_of_eq ?_) simp only [map_one, mul_one, MonoidHom.one_apply] +@[to_additive] +lemma TopologicalGroup.isInducing_iff_nhds_one + {H : Type*} [Group H] [TopologicalSpace H] [TopologicalGroup H] {f : G →* H} : + Topology.IsInducing f ↔ 𝓝 (1 : G) = (𝓝 (1 : H)).comap f := by + rw [Topology.isInducing_iff_nhds] + refine ⟨(f.map_one ▸ · 1), fun hf x ↦ ?_⟩ + rw [← nhds_translation_mul_inv, ← nhds_translation_mul_inv (f x), Filter.comap_comap, hf, + Filter.comap_comap] + congr 1 + ext; simp + -- TODO: unify with `QuotientGroup.isOpenQuotientMap_mk` /-- Let `A` and `B` be topological groups, and let `φ : A → B` be a continuous surjective group homomorphism. Assume furthermore that `φ` is a quotient map (i.e., `V ⊆ B` From bb6faa6b4a4b0d20736c8d967444b7affb088987 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Mon, 30 Dec 2024 15:46:54 +0000 Subject: [PATCH 037/189] feat(AlgebraicGeometry): the small affine Zariski site (#19558) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define the small affine Zariski site of `X` (`X.AffineZariskiSite`), whose elements are affine open sets of `X`, and whose arrows are basic open sets `D(f) ⟶ U` for any `f : Γ(X, U)`. This is a dense subsite of `X.Opens` (with respect to `Opens.grothendieckTopology X`) via the inclusion functor, which gives an equivalence of categories of sheaves (`sheafEquiv`). --- Mathlib.lean | 1 + .../Sites/SmallAffineZariski.lean | 192 ++++++++++++++++++ 2 files changed, 193 insertions(+) create mode 100644 Mathlib/AlgebraicGeometry/Sites/SmallAffineZariski.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1625c95f15a04..8b6803e4014ff 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1029,6 +1029,7 @@ import Mathlib.AlgebraicGeometry.Sites.BigZariski import Mathlib.AlgebraicGeometry.Sites.Etale import Mathlib.AlgebraicGeometry.Sites.MorphismProperty import Mathlib.AlgebraicGeometry.Sites.Small +import Mathlib.AlgebraicGeometry.Sites.SmallAffineZariski import Mathlib.AlgebraicGeometry.Spec import Mathlib.AlgebraicGeometry.SpreadingOut import Mathlib.AlgebraicGeometry.Stalk diff --git a/Mathlib/AlgebraicGeometry/Sites/SmallAffineZariski.lean b/Mathlib/AlgebraicGeometry/Sites/SmallAffineZariski.lean new file mode 100644 index 0000000000000..04ec6147501ff --- /dev/null +++ b/Mathlib/AlgebraicGeometry/Sites/SmallAffineZariski.lean @@ -0,0 +1,192 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.AlgebraicGeometry.AffineScheme +import Mathlib.CategoryTheory.Sites.DenseSubsite.InducedTopology + +/-! + +# The small affine Zariski site + +`X.AffineZariskiSite` is the small affine Zariski site of `X`, whose elements are affine open +sets of `X`, and whose arrows are basic open sets `D(f) ⟶ U` for any `f : Γ(X, U)`. + +Every presieve on `U` is then given by a `Set Γ(X, U)` (`presieveOfSections_surjective`), and +we endow `X.AffineZariskiSite` with `grothendieckTopology X`, such that `s : Set Γ(X, U)` is +a cover if and only if `Ideal.span s = ⊤` (`generate_presieveOfSections_mem_grothendieckTopology`). + +This is a dense subsite of `X.Opens` (with respect to `Opens.grothendieckTopology X`) via the +inclusion functor `toOpensFunctor X`, +which gives an equivalence of categories of sheaves (`sheafEquiv`). + +Note that this differs from the definition on stacks project where the arrows in the small affine +Zariski site are arbitrary inclusions. + +-/ + +universe u + +open CategoryTheory + +noncomputable section + +namespace AlgebraicGeometry + +variable {X : Scheme.{u}} + +/-- +`X.AffineZariskiSite` is the small affine Zariski site of `X`, whose elements are affine open +sets of `X`, and whose arrows are basic open sets `D(f) ⟶ U` for any `f : Γ(X, U)`. + +Note that this differs from the definition on stacks project where the arrows in the small affine +Zariski site are arbitrary inclusions. +-/ +def Scheme.AffineZariskiSite (X : Scheme.{u}) : Type u := { U : X.Opens // IsAffineOpen U } + +namespace Scheme.AffineZariskiSite + +/-- The inclusion from `X.AffineZariskiSite` to `X.Opens`. -/ +abbrev toOpens (U : X.AffineZariskiSite) : X.Opens := U.1 + +instance : Preorder X.AffineZariskiSite where + le U V := ∃ f : Γ(X, V.toOpens), X.basicOpen f = U.toOpens + le_refl U := ⟨1, Scheme.basicOpen_of_isUnit _ isUnit_one⟩ + le_trans := by + rintro ⟨U, hU⟩ ⟨V, hV⟩ ⟨W, hW⟩ ⟨f, rfl⟩ ⟨g, rfl⟩ + exact hW.basicOpen_basicOpen_is_basicOpen g f + +lemma toOpens_mono : + Monotone (toOpens (X := X)) := by + rintro ⟨U, hU⟩ ⟨V, hV⟩ ⟨f, rfl⟩ + exact X.basicOpen_le _ + +lemma toOpens_injective : Function.Injective (toOpens (X := X)) := Subtype.val_injective + +instance : PartialOrder X.AffineZariskiSite where + le_antisymm _ _ hUV hVU := Subtype.ext ((toOpens_mono hUV).antisymm (toOpens_mono hVU)) + +/-- The basic open set of a section, as an element of `AffineZariskiSite`. -/ +def basicOpen (U : X.AffineZariskiSite) (f : Γ(X, U.toOpens)) : X.AffineZariskiSite := + ⟨X.basicOpen f, U.2.basicOpen f⟩ + +lemma basicOpen_le (U : X.AffineZariskiSite) (f : Γ(X, U.toOpens)) : U.basicOpen f ≤ U := + ⟨f, rfl⟩ + +variable (X) in +/-- The inclusion functor from `X.AffineZariskiSite` to `X.Opens`. -/ +def toOpensFunctor : X.AffineZariskiSite ⥤ X.Opens := toOpens_mono.functor + +instance : (toOpensFunctor X).Faithful where + +instance : (toOpensFunctor X).IsLocallyFull (Opens.grothendieckTopology X) where + functorPushforward_imageSieve_mem := by + intro U V h x hx + obtain ⟨f, hfU, hxf⟩ := V.2.exists_basicOpen_le ⟨x, hx⟩ (h.le hx) + exact ⟨X.basicOpen f, homOfLE hfU, ⟨V.basicOpen f, + ⟨_, (X.basicOpen_res f h.op).trans (inf_eq_right.mpr hfU)⟩, 𝟙 _, + ⟨⟨f, rfl⟩, rfl⟩, rfl⟩, hxf⟩ + +instance : (toOpensFunctor X).IsCoverDense (Opens.grothendieckTopology X) where + is_cover := by + intros U x hx + obtain ⟨_, ⟨V, hV, rfl⟩, hxV, hVU⟩ := (isBasis_affine_open X).exists_subset_of_mem_open hx U.2 + exact ⟨V, homOfLE hVU, ⟨⟨V, hV⟩, 𝟙 _, homOfLE hVU, rfl⟩, hxV⟩ + +variable (X) in +/-- The grothendieck topology on `X.AffineZariskiSite` induced from the topology on `X.Opens`. +Also see `mem_grothendieckTopology_iff_sectionsOfPresieve`. -/ +def grothendieckTopology : GrothendieckTopology X.AffineZariskiSite := + (toOpensFunctor X).inducedTopology (Opens.grothendieckTopology X) + +lemma mem_grothendieckTopology {U : X.AffineZariskiSite} {S : Sieve U} : + S ∈ grothendieckTopology X U ↔ + ∀ x ∈ U.toOpens, ∃ (V : _) (f : V ⟶ U), S.arrows f ∧ x ∈ V.toOpens := by + apply forall₂_congr fun x hxU ↦ ⟨?_, ?_⟩ + · rintro ⟨V, f, ⟨W, g, h, hg, rfl⟩, hxV⟩ + exact ⟨W, g, hg, h.le hxV⟩ + · rintro ⟨W, g, hg, hxW⟩ + exact ⟨W.toOpens, homOfLE (toOpens_mono g.le), ⟨W, g, 𝟙 _, hg, rfl⟩, hxW⟩ + +instance : (toOpensFunctor X).IsDenseSubsite + (grothendieckTopology X) (Opens.grothendieckTopology X) where + functorPushforward_mem_iff := Iff.rfl + +/-- The presieve associated to a set of sections. +This is a surjection, see `presieveOfSections_surjective`. -/ +def presieveOfSections (U : X.AffineZariskiSite) (s : Set Γ(X, U.toOpens)) : Presieve U := + fun V _ ↦ ∃ f ∈ s, X.basicOpen f = V.toOpens + +/-- The set of sections associated to a presieve. -/ +def sectionsOfPresieve {U : X.AffineZariskiSite} (P : Presieve U) : Set Γ(X, U.toOpens) := + { f | P (homOfLE (U.basicOpen_le f)) } + +lemma presieveOfSections_sectionsOfPresieve {U : X.AffineZariskiSite} (P : Presieve U) : + presieveOfSections U (sectionsOfPresieve P) = P := by + refine funext₂ fun ⟨V, hV⟩ ⟨f, hf⟩ ↦ eq_iff_iff.mpr ⟨?_, ?_⟩ + · rintro ⟨_, H, rfl⟩ + exact H + · intro H + obtain rfl : _ = V := hf + exact ⟨_, H, rfl⟩ + +lemma presieveOfSections_surjective {U : X.AffineZariskiSite} : + Function.Surjective (presieveOfSections U) := + fun _ ↦ ⟨_, presieveOfSections_sectionsOfPresieve _⟩ + +lemma presieveOfSections_eq_ofArrows (U : X.AffineZariskiSite) (s : Set Γ(X, U.toOpens)) : + presieveOfSections U s = .ofArrows _ (fun i : s ↦ homOfLE (U.basicOpen_le i.1)) := by + refine funext₂ fun ⟨V, hV⟩ ⟨f, hf⟩ ↦ eq_iff_iff.mpr ⟨?_, ?_⟩ + · rintro ⟨f, hfs, rfl⟩ + exact .mk (ι := s) ⟨f, hfs⟩ + · rintro ⟨⟨f, hfs⟩⟩ + exact ⟨f, hfs, rfl⟩ + +lemma generate_presieveOfSections + {U V : X.AffineZariskiSite} {s : Set Γ(X, U.toOpens)} {f : V ⟶ U} : + Sieve.generate (presieveOfSections U s) f ↔ ∃ f ∈ s, ∃ g, X.basicOpen (f * g) = V.toOpens := by + obtain ⟨V, hV⟩ := V + constructor + · rintro ⟨⟨W, hW⟩, ⟨f₁, hf₁⟩, -, ⟨f₂, hf₂s, rfl⟩, rfl⟩ + subst hf₁ + obtain ⟨f₃, hf₃⟩ := U.2.basicOpen_basicOpen_is_basicOpen f₂ f₁ + refine ⟨f₂, hf₂s, f₃, ?_⟩ + rw [X.basicOpen_mul, hf₃, inf_eq_right] + exact X.basicOpen_le _ + · rintro ⟨f₁, hf₁s, f₂, rfl⟩ + refine ⟨U.basicOpen f₁, ⟨f₂ |_ᵣ _, ?_⟩, ⟨f₁, rfl⟩, ⟨f₁, hf₁s, rfl⟩, rfl⟩ + exact (X.basicOpen_res _ _).trans (X.basicOpen_mul _ _).symm + +lemma generate_presieveOfSections_mem_grothendieckTopology + {U : X.AffineZariskiSite} {s : Set Γ(X, U.toOpens)} : + Sieve.generate (presieveOfSections U s) ∈ grothendieckTopology X U ↔ Ideal.span s = ⊤ := by + rw [← U.2.self_le_basicOpen_union_iff, mem_grothendieckTopology, SetLike.le_def] + refine forall₂_congr fun x hx ↦ ?_ + simp only [exists_and_left, TopologicalSpace.Opens.iSup_mk, + TopologicalSpace.Opens.carrier_eq_coe, Set.iUnion_coe_set, TopologicalSpace.Opens.mem_mk, + Set.mem_iUnion, SetLike.mem_coe, exists_prop, generate_presieveOfSections] + constructor + · simp only [basicOpen_mul] + rintro ⟨⟨V, hV⟩, ⟨f, hfs, g, rfl⟩, -, hxV⟩ + exact ⟨f, hfs, hxV.1⟩ + · rintro ⟨f, hfs, hxf⟩ + refine ⟨U.basicOpen _, ⟨f, hfs, 1, rfl⟩, ⟨_, rfl⟩, by simpa using hxf⟩ + +lemma mem_grothendieckTopology_iff_sectionsOfPresieve + {U : X.AffineZariskiSite} {S : Sieve U} : + S ∈ grothendieckTopology X U ↔ Ideal.span (sectionsOfPresieve S.1) = ⊤ := by + rw [← generate_presieveOfSections_mem_grothendieckTopology, presieveOfSections_sectionsOfPresieve, + Sieve.generate_sieve] + +variable {A} [Category A] +variable [∀ (U : X.Opensᵒᵖ), Limits.HasLimitsOfShape (StructuredArrow U (toOpensFunctor X).op) A] + +/-- The category of sheaves on `X.AffineZariskiSite` is equivalent to the categories of sheaves +over `X`. -/ +abbrev sheafEquiv : Sheaf (grothendieckTopology X) A ≌ TopCat.Sheaf A X := + (toOpensFunctor X).sheafInducedTopologyEquivOfIsCoverDense _ _ + +end Scheme.AffineZariskiSite + +end AlgebraicGeometry From 9df7ff0b69d673339d11e859ab526223ccc63742 Mon Sep 17 00:00:00 2001 From: smorel394 Date: Mon, 30 Dec 2024 18:55:52 +0000 Subject: [PATCH 038/189] feat(CategoryTheory/Triangulated/Opposite): the opposite of a triangulated functor (#20278) A functor between pretriangulated categories is a triangulated functor if and only if its opposite is a triangulated functor. Co-authored-by: smorel394 <67864981+smorel394@users.noreply.github.com> --- Mathlib.lean | 1 + .../Triangulated/Opposite/Functor.lean | 221 ++++++++++++++++++ 2 files changed, 222 insertions(+) create mode 100644 Mathlib/CategoryTheory/Triangulated/Opposite/Functor.lean diff --git a/Mathlib.lean b/Mathlib.lean index 8b6803e4014ff..9b2281b6fb61a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2207,6 +2207,7 @@ import Mathlib.CategoryTheory.Triangulated.Basic import Mathlib.CategoryTheory.Triangulated.Functor import Mathlib.CategoryTheory.Triangulated.HomologicalFunctor import Mathlib.CategoryTheory.Triangulated.Opposite.Basic +import Mathlib.CategoryTheory.Triangulated.Opposite.Functor import Mathlib.CategoryTheory.Triangulated.Opposite.Pretriangulated import Mathlib.CategoryTheory.Triangulated.Opposite.Triangle import Mathlib.CategoryTheory.Triangulated.Pretriangulated diff --git a/Mathlib/CategoryTheory/Triangulated/Opposite/Functor.lean b/Mathlib/CategoryTheory/Triangulated/Opposite/Functor.lean new file mode 100644 index 0000000000000..16109ed271f3f --- /dev/null +++ b/Mathlib/CategoryTheory/Triangulated/Opposite/Functor.lean @@ -0,0 +1,221 @@ +/- +Copyright (c) 2024 Sophie Morel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sophie Morel +-/ +import Mathlib.CategoryTheory.Triangulated.Opposite.Pretriangulated + +/-! +# Opposites of functors between pretriangulated categories, + +If `F : C ⥤ D` is a functor between pretriangulated categories, we prove that +`F` is a triangulated functor if and only if `F.op` is a triangulated functor. +In order to do this, we first show that a `CommShift` structure on `F` naturally +gives one on `F.op` (for the shifts on `Cᵒᵖ` and `Dᵒᵖ` defined in +`CategoryTheory.Triangulated.Opposite.Basic`), and we then prove +that `F.mapTriangle.op` and `F.op.mapTriangle` correspond to each other via the +equivalences `(Triangle C)ᵒᵖ ≌ Triangle Cᵒᵖ` and `(Triangle D)ᵒᵖ ≌ Triangle Dᵒᵖ` +given by `CategoryTheory.Pretriangulated.triangleOpEquivalence`. + +-/ + +namespace CategoryTheory + +variable {C D : Type*} [Category C] [Category D] [HasShift C ℤ] [HasShift D ℤ] (F : C ⥤ D) + [F.CommShift ℤ] + +namespace Pretriangulated.Opposite + +/-- If `F` commutes with shifts, so does `F.op`, for the shifts chosen on `Cᵒᵖ` in +`CategoryTheory.Triangulated.Opposite.Basic`. +-/ +noncomputable scoped instance commShiftOpInt : F.op.CommShift ℤ := by + letI F' : OppositeShift C ℤ ⥤ OppositeShift D ℤ := F.op + letI : F'.CommShift ℤ := F.commShiftOp ℤ + apply F'.commShiftPullback + +end Pretriangulated.Opposite + +namespace Functor + +open Category Limits Pretriangulated Opposite + +@[reassoc] +lemma op_commShiftIso_hom_app (X : Cᵒᵖ) (n m : ℤ) (h : n + m = 0): + (F.op.commShiftIso n).hom.app X = + (F.map ((shiftFunctorOpIso C n m h).hom.app X).unop).op ≫ + ((F.commShiftIso m).inv.app X.unop).op ≫ + (shiftFunctorOpIso D n m h).inv.app (op (F.obj X.unop)) := by + obtain rfl : m = -n := by omega + rfl + +@[reassoc] +lemma op_commShiftIso_inv_app (X : Cᵒᵖ) (n m : ℤ) (h : n + m = 0): + (F.op.commShiftIso n).inv.app X = + (shiftFunctorOpIso D n m h).hom.app (op (F.obj X.unop)) ≫ + ((F.commShiftIso m).hom.app X.unop).op ≫ + (F.map ((shiftFunctorOpIso C n m h).inv.app X).unop).op := by + rw [← cancel_epi ((F.op.commShiftIso n).hom.app X), Iso.hom_inv_id_app, + op_commShiftIso_hom_app _ X n m h, assoc, assoc] + simp [← op_comp, ← F.map_comp] + +@[reassoc] +lemma shift_map_op {X Y : C} (f : X ⟶ Y) (n : ℤ) : + (F.map f).op⟦n⟧' = (F.op.commShiftIso n).inv.app _ ≫ + (F.map (f.op⟦n⟧').unop).op ≫ (F.op.commShiftIso n).hom.app _ := + (NatIso.naturality_1 (F.op.commShiftIso n) f.op).symm + +@[reassoc] +lemma map_shift_unop {X Y : Cᵒᵖ} (f : X ⟶ Y) (n : ℤ) : + F.map ((f⟦n⟧').unop) = ((F.op.commShiftIso n).inv.app Y).unop ≫ + ((F.map f.unop).op⟦n⟧').unop ≫ ((F.op.commShiftIso n).hom.app X).unop := by + simp [shift_map_op] + +@[reassoc] +lemma map_opShiftFunctorEquivalence_unitIso_hom_app_unop (X : Cᵒᵖ) (n : ℤ) : + F.map ((opShiftFunctorEquivalence C n).unitIso.hom.app X).unop = + (F.commShiftIso n).hom.app _ ≫ + (((F.op).commShiftIso n).inv.app X).unop⟦n⟧' ≫ + ((opShiftFunctorEquivalence D n).unitIso.hom.app (op _)).unop := by + dsimp [opShiftFunctorEquivalence] + simp only [map_comp, unop_comp, Quiver.Hom.unop_op, assoc, + map_shiftFunctorCompIsoId_hom_app, commShiftIso_hom_naturality_assoc, + op_commShiftIso_inv_app _ _ _ _ (add_neg_cancel n)] + congr 3 + rw [← Functor.map_comp_assoc, ← unop_comp, + Iso.inv_hom_id_app] + dsimp + rw [map_id, id_comp] + +@[reassoc] +lemma map_opShiftFunctorEquivalence_unitIso_inv_app_unop (X : Cᵒᵖ) (n : ℤ) : + F.map ((opShiftFunctorEquivalence C n).unitIso.inv.app X).unop = + ((opShiftFunctorEquivalence D n).unitIso.inv.app (op (F.obj X.unop))).unop ≫ + (((F.op).commShiftIso n).hom.app X).unop⟦n⟧' ≫ + ((F.commShiftIso n).inv.app _) := by + rw [← cancel_mono (F.map ((opShiftFunctorEquivalence C n).unitIso.hom.app X).unop), + ← F.map_comp, ← unop_comp, Iso.hom_inv_id_app, + map_opShiftFunctorEquivalence_unitIso_hom_app_unop, assoc, assoc, + Iso.inv_hom_id_app_assoc, ← Functor.map_comp_assoc, ← unop_comp] + simp + +@[reassoc] +lemma map_opShiftFunctorEquivalence_counitIso_hom_app_unop (X : Cᵒᵖ) (n : ℤ) : + F.map ((opShiftFunctorEquivalence C n).counitIso.hom.app X).unop = + ((opShiftFunctorEquivalence D n).counitIso.hom.app (op (F.obj X.unop))).unop ≫ + (((F.commShiftIso n).inv.app X.unop).op⟦n⟧').unop ≫ + ((F.op.commShiftIso n).hom.app (op (X.unop⟦n⟧))).unop := by + apply Quiver.Hom.op_inj + dsimp [opShiftFunctorEquivalence] + rw [assoc, F.op_commShiftIso_hom_app_assoc _ _ _ (add_neg_cancel n), map_comp, + map_shiftFunctorCompIsoId_inv_app_assoc, op_comp, op_comp_assoc, op_comp_assoc, + NatTrans.naturality_assoc, op_map, Iso.inv_hom_id_app_assoc, Quiver.Hom.unop_op] + +@[reassoc] +lemma map_opShiftFunctorEquivalence_counitIso_inv_app_unop (X : Cᵒᵖ) (n : ℤ) : + F.map ((opShiftFunctorEquivalence C n).counitIso.inv.app X).unop = + ((F.op.commShiftIso n).inv.app (op (X.unop⟦n⟧))).unop ≫ + (((F.commShiftIso n).hom.app X.unop).op⟦n⟧').unop ≫ + ((opShiftFunctorEquivalence D n).counitIso.inv.app (op (F.obj X.unop))).unop := by + rw [← cancel_epi (F.map ((opShiftFunctorEquivalence C n).counitIso.hom.app X).unop), + ← F.map_comp, ← unop_comp, Iso.inv_hom_id_app, + map_opShiftFunctorEquivalence_counitIso_hom_app_unop] + dsimp + simp only [map_id, assoc, Iso.unop_hom_inv_id_app_assoc, ← Functor.map_comp_assoc, + ← unop_comp, Iso.inv_hom_id_app_assoc, ← unop_comp_assoc, ← op_comp, + Iso.inv_hom_id_app] + simp + +variable [HasZeroObject C] [Preadditive C] [∀ (n : ℤ), (shiftFunctor C n).Additive] + [Pretriangulated C] [HasZeroObject D] [Preadditive D] + [∀ (n : ℤ), (shiftFunctor D n).Additive] [Pretriangulated D] + +/-- +If `F : C ⥤ D` commutes with shifts, this expresses the compatibility of `F.mapTriangle` +with the equivalences `Pretriangulated.triangleOpEquivalence` on `C` and `D`. +-/ +@[simps!] +noncomputable def mapTriangleOpCompTriangleOpEquivalenceFunctorApp (T : Triangle C) : + (triangleOpEquivalence D).functor.obj (op (F.mapTriangle.obj T)) ≅ + F.op.mapTriangle.obj ((triangleOpEquivalence C).functor.obj (op T)) := + Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) + (by dsimp; simp) (by dsimp; simp) (by + dsimp + simp only [map_comp, shift_map_op, map_id, comp_id, op_comp, op_unop, + map_opShiftFunctorEquivalence_counitIso_inv_app_unop, + opShiftFunctorEquivalence_inverse, opShiftFunctorEquivalence_functor, + Quiver.Hom.op_unop, assoc, id_comp]) + +/-- +If `F : C ⥤ D` commutes with shifts, this expresses the compatibility of `F.mapTriangle` +with the equivalences `Pretriangulated.triangleOpEquivalence` on `C` and `D`. +-/ +noncomputable def mapTriangleOpCompTriangleOpEquivalenceFunctor : + F.mapTriangle.op ⋙ (triangleOpEquivalence D).functor ≅ + (triangleOpEquivalence C).functor ⋙ F.op.mapTriangle := + NatIso.ofComponents + (fun T ↦ F.mapTriangleOpCompTriangleOpEquivalenceFunctorApp T.unop) + (by intros; ext <;> dsimp <;> simp only [comp_id, id_comp]) + +/-- +If `F : C ⥤ D` commutes with shifts, this is the 2-commutative square of categories +`CategoryTheory.Functor.mapTriangleOpCompTriangleOpEquivalenceFunctor`. +-/ +noncomputable instance : + CatCommSq (F.mapTriangle.op) (triangleOpEquivalence C).functor + (triangleOpEquivalence D).functor F.op.mapTriangle := + ⟨F.mapTriangleOpCompTriangleOpEquivalenceFunctor⟩ + +/-- +Vertical inverse of the 2-commutative square of +`CategoryTheory.Functor.mapTriangleOpCompTriangleOpEquivalenceFunctor`. +-/ +noncomputable instance : + CatCommSq (F.op.mapTriangle) (triangleOpEquivalence C).inverse + (triangleOpEquivalence D).inverse F.mapTriangle.op := + CatCommSq.vInv (F.mapTriangle.op) (triangleOpEquivalence C) + (triangleOpEquivalence D) F.op.mapTriangle inferInstance + +/-- +If `F : C ⥤ D` commutes with shifts, this expresses the compatibility of `F.mapTriangle` +with the equivalences `Pretriangulated.triangleOpEquivalence` on `C` and `D`. +-/ +noncomputable def opMapTriangleCompTriangleOpEquivalenceInverse : + F.op.mapTriangle ⋙ (triangleOpEquivalence D).inverse ≅ + (triangleOpEquivalence C).inverse ⋙ F.mapTriangle.op := + CatCommSq.iso (F.op.mapTriangle) (triangleOpEquivalence C).inverse + (triangleOpEquivalence D).inverse F.mapTriangle.op + +/-- If `F` is triangulated, so is `F.op`. +-/ +lemma isTriangulated_op [F.IsTriangulated] : F.op.IsTriangulated where + map_distinguished T dT := by + rw [mem_distTriang_op_iff] + exact Pretriangulated.isomorphic_distinguished _ + ((F.map_distinguished _ (unop_distinguished _ dT))) _ + (((opMapTriangleCompTriangleOpEquivalenceInverse F).symm.app T).unop) + +/-- If `F.op` is triangulated, so is `F`. +-/ +lemma isTriangulated_of_op [F.op.IsTriangulated] : F.IsTriangulated where + map_distinguished T dT := by + have := distinguished_iff_of_iso ((triangleOpEquivalence D).unitIso.app + (Opposite.op (F.mapTriangle.obj T))).unop + rw [Functor.id_obj, Opposite.unop_op (F.mapTriangle.obj T)] at this + rw [← this, Functor.comp_obj, ← mem_distTriang_op_iff, ← Functor.op_obj, ← Functor.comp_obj, + distinguished_iff_of_iso ((mapTriangleOpCompTriangleOpEquivalenceFunctor F).app + (Opposite.op T))] + apply F.op.map_distinguished + have := distinguished_iff_of_iso ((triangleOpEquivalence C).unitIso.app (Opposite.op T)).unop + rw [Functor.id_obj, Opposite.unop_op T] at this + rw [← this, Functor.comp_obj, ← mem_distTriang_op_iff] at dT + exact dT + +/-- `F` is triangulated if and only if `F.op` is triangulated. +-/ +lemma op_isTriangulated_iff : F.op.IsTriangulated ↔ F.IsTriangulated := + ⟨fun _ ↦ F.isTriangulated_of_op, fun _ ↦ F.isTriangulated_op⟩ + +end Functor + +end CategoryTheory From 3616fa83ff48ad8fe1dcc84f0f4c899abe86abcf Mon Sep 17 00:00:00 2001 From: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Date: Mon, 30 Dec 2024 19:05:26 +0000 Subject: [PATCH 039/189] chore(Denumerable): simplify proof of ofNat_surjective (#20344) No need for auxiliary theorem, and a bit of simplification. --- Mathlib/Logic/Denumerable.lean | 24 +++++++----------------- 1 file changed, 7 insertions(+), 17 deletions(-) diff --git a/Mathlib/Logic/Denumerable.lean b/Mathlib/Logic/Denumerable.lean index d6160e1858ab6..b45cea7453b1a 100644 --- a/Mathlib/Logic/Denumerable.lean +++ b/Mathlib/Logic/Denumerable.lean @@ -234,32 +234,22 @@ def ofNat (s : Set ℕ) [DecidablePred (· ∈ s)] [Infinite s] : ℕ → s | 0 => ⊥ | n + 1 => succ (ofNat s n) -theorem ofNat_surjective_aux : ∀ {x : ℕ} (hx : x ∈ s), ∃ n, ofNat s n = ⟨x, hx⟩ - | x => fun hx => by +theorem ofNat_surjective : Surjective (ofNat s) + | ⟨x, hx⟩ => by set t : List s := ((List.range x).filter fun y => y ∈ s).pmap (fun (y : ℕ) (hy : y ∈ s) => ⟨y, hy⟩) (by intros a ha; simpa using (List.mem_filter.mp ha).2) with ht have hmt : ∀ {y : s}, y ∈ t ↔ y < ⟨x, hx⟩ := by simp [List.mem_filter, Subtype.ext_iff_val, ht] - have wf : ∀ m : s, List.maximum t = m → ↑m < x := fun m hmax => by - simpa using hmt.mp (List.maximum_mem hmax) cases' hmax : List.maximum t with m · refine ⟨0, le_antisymm bot_le (le_of_not_gt fun h => List.not_mem_nil (⊥ : s) ?_)⟩ rwa [← List.maximum_eq_bot.1 hmax, hmt] - cases' ofNat_surjective_aux m.2 with a ha - refine ⟨a + 1, le_antisymm ?_ ?_⟩ <;> rw [ofNat] - · refine succ_le_of_lt ?_ - rw [ha] - exact wf _ hmax - · refine le_succ_of_forall_lt_le fun z hz => ?_ - rw [ha] - cases m - exact List.le_maximum_of_mem (hmt.2 hz) hmax -decreasing_by - tauto - -theorem ofNat_surjective : Surjective (ofNat s) := fun ⟨_, hx⟩ => ofNat_surjective_aux hx + have wf : ↑m < x := by simpa using hmt.mp (List.maximum_mem hmax) + rcases ofNat_surjective m with ⟨a, rfl⟩ + refine ⟨a + 1, le_antisymm (succ_le_of_lt wf) ?_⟩ + exact le_succ_of_forall_lt_le fun z hz => List.le_maximum_of_mem (hmt.2 hz) hmax + termination_by n => n.val @[simp] theorem ofNat_range : Set.range (ofNat s) = Set.univ := From db269ba2df60d788b1321d1f02010e991199e849 Mon Sep 17 00:00:00 2001 From: Sidharth Hariharan Date: Mon, 30 Dec 2024 19:14:44 +0000 Subject: [PATCH 040/189] feat(NumberTheory/Divisors): proving that n has at most n divisors. (#20312) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR contains a proof that for some `n : ℕ`, `n.divisors` has cardinality at most `n`. It is useful in a number of scenarios, for instance when one wishes to construct upper-bounds for sums over the set of divisors of `n`. Co-authored-by: Sidharth Hariharan --- Mathlib/NumberTheory/Divisors.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index e3d2f24689a32..e0af39d898670 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -133,6 +133,12 @@ theorem divisor_le {m : ℕ} : n ∈ divisors m → n ≤ m := by theorem divisors_subset_of_dvd {m : ℕ} (hzero : n ≠ 0) (h : m ∣ n) : divisors m ⊆ divisors n := Finset.subset_iff.2 fun _x hx => Nat.mem_divisors.mpr ⟨(Nat.mem_divisors.mp hx).1.trans h, hzero⟩ +theorem card_divisors_le_self (n : ℕ) : #n.divisors ≤ n := calc + _ ≤ #(Ico 1 (n + 1)) := by + apply card_le_card + simp only [divisors, filter_subset] + _ = n := by rw [card_Ico, add_tsub_cancel_right] + theorem divisors_subset_properDivisors {m : ℕ} (hzero : n ≠ 0) (h : m ∣ n) (hdiff : m ≠ n) : divisors m ⊆ properDivisors n := by apply Finset.subset_iff.2 From 1dce20e08835d80c42f93b049457834001c6f309 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 30 Dec 2024 19:50:57 +0000 Subject: [PATCH 041/189] chore: Replace `open scoped Classical` with `open scoped Classical in` or `classical` (#20325) I haven't done all the files, but a fair number of them. The only changes in this PR are: * Replace `open scoped Classical` with `open scoped Classical in` (when needed for the statement) or `classical` in proofs. * Some term mode proofs with `if` have been converted to tactic mode proofs with `by_cases`. * Some adjacent `open ...` statements have been merged into a single `open ...` --- Mathlib/Analysis/Analytic/CPolynomial.lean | 2 +- Mathlib/Analysis/Analytic/Constructions.lean | 4 ++- Mathlib/Analysis/Analytic/IsolatedZeros.lean | 4 +-- Mathlib/Analysis/Asymptotics/Asymptotics.lean | 8 +++-- Mathlib/Analysis/BoxIntegral/Box/Basic.lean | 3 +- .../BoxIntegral/Partition/Additive.lean | 2 +- .../Analysis/BoxIntegral/Partition/Basic.lean | 30 +++++++++++------ .../BoxIntegral/Partition/Filter.lean | 3 +- .../Analysis/BoxIntegral/Partition/Split.lean | 17 ++++++++-- .../Partition/SubboxInduction.lean | 2 +- .../BoxIntegral/Partition/Tagged.lean | 32 +++++++++++-------- Mathlib/Analysis/Calculus/ContDiff/Defs.lean | 1 - Mathlib/Analysis/Calculus/Darboux.lean | 2 +- Mathlib/Analysis/Calculus/Deriv/Add.lean | 1 - Mathlib/Analysis/Calculus/Deriv/Comp.lean | 2 +- Mathlib/Analysis/Calculus/Deriv/Prod.lean | 5 +-- Mathlib/Analysis/Calculus/Deriv/ZPow.lean | 5 +-- Mathlib/Analysis/Calculus/FDeriv/Add.lean | 21 ++++++------ Mathlib/Analysis/Calculus/FDeriv/Basic.lean | 9 +++--- Mathlib/Analysis/Calculus/FDeriv/Comp.lean | 5 +-- Mathlib/Analysis/Calculus/FDeriv/Equiv.lean | 5 +-- Mathlib/Analysis/Calculus/FDeriv/Prod.lean | 5 +-- .../Calculus/FDeriv/RestrictScalars.lean | 5 +-- Mathlib/Analysis/Calculus/FDeriv/Star.lean | 2 -- Mathlib/Analysis/Calculus/LogDeriv.lean | 2 +- Mathlib/Analysis/Convex/Combination.lean | 18 ++++++++--- Mathlib/Analysis/Convex/Cone/Basic.lean | 5 +-- Mathlib/Analysis/Convex/Cone/InnerDual.lean | 5 +-- Mathlib/Analysis/Convex/Exposed.lean | 7 ++-- Mathlib/Analysis/Convex/Extrema.lean | 5 +-- Mathlib/Analysis/Convex/Extreme.lean | 6 ++-- Mathlib/Analysis/Convex/Function.lean | 2 -- Mathlib/Analysis/Convex/Independent.lean | 6 ++-- Mathlib/Analysis/Convex/Jensen.lean | 5 +-- Mathlib/Analysis/Convex/KreinMilman.lean | 1 - Mathlib/Analysis/InnerProductSpace/Dual.lean | 1 - .../InnerProductSpace/LinearPMap.lean | 5 ++- Mathlib/Analysis/Normed/Module/Dual.lean | 1 - Mathlib/Analysis/Normed/Operator/Banach.lean | 1 - Mathlib/Analysis/Normed/Ring/Units.lean | 2 -- Mathlib/Analysis/RCLike/Lemmas.lean | 2 -- .../SpecialFunctions/Pow/Asymptotics.lean | 1 - .../SpecialFunctions/Pow/Complex.lean | 1 - .../SpecialFunctions/Pow/Continuity.lean | 5 +-- .../Analysis/SpecialFunctions/Pow/NNReal.lean | 8 ++--- .../Analysis/SpecialFunctions/Pow/Real.lean | 5 +-- .../SpecialFunctions/Trigonometric/Basic.lean | 1 - .../Trigonometric/Inverse.lean | 7 +--- Mathlib/Analysis/SpecificLimits/Basic.lean | 8 ++--- Mathlib/CategoryTheory/Countable.lean | 2 -- .../CategoryTheory/FinCategory/AsType.lean | 2 -- Mathlib/CategoryTheory/FinCategory/Basic.lean | 3 +- Mathlib/CategoryTheory/FintypeCat.lean | 3 -- Mathlib/CategoryTheory/GlueData.lean | 5 +-- Mathlib/CategoryTheory/Limits/Bicones.lean | 6 ++-- .../Limits/Shapes/Biproducts.lean | 31 +++++++++++++----- .../Limits/Shapes/FiniteProducts.lean | 2 -- .../Limits/Shapes/ZeroMorphisms.lean | 10 +++--- .../CategoryTheory/Monoidal/Preadditive.lean | 23 ++++++++++--- .../Preadditive/Biproducts.lean | 12 +++++-- .../Preadditive/HomOrthogonal.lean | 5 ++- Mathlib/CategoryTheory/Preadditive/Schur.lean | 3 +- Mathlib/CategoryTheory/Sites/Canonical.lean | 1 - Mathlib/Combinatorics/HalesJewett.lean | 3 +- .../SimpleGraph/Regularity/Uniform.lean | 3 +- Mathlib/Data/Finite/Card.lean | 6 ++-- Mathlib/Data/Finite/Prod.lean | 6 ++-- Mathlib/Data/Finset/Basic.lean | 2 -- Mathlib/Data/Finset/Finsupp.lean | 7 ++-- Mathlib/Data/Finset/Functor.lean | 5 +-- Mathlib/Data/Finsupp/Interval.lean | 14 +++++--- Mathlib/Data/Fintype/Card.lean | 3 +- Mathlib/Data/Fintype/Order.lean | 3 +- Mathlib/Data/Fintype/Prod.lean | 8 ++--- Mathlib/Data/Fintype/Sum.lean | 2 -- Mathlib/Data/Real/Archimedean.lean | 5 ++- Mathlib/Data/Real/Basic.lean | 3 +- Mathlib/Data/Real/ENatENNReal.lean | 1 - Mathlib/Data/Real/Hyperreal.lean | 11 ++++--- Mathlib/Data/Set/BoolIndicator.lean | 3 +- Mathlib/Data/Set/Countable.lean | 1 - Mathlib/Data/Set/Finite/Basic.lean | 5 +-- Mathlib/Data/Set/Finite/Lattice.lean | 2 -- Mathlib/Data/Set/Finite/Monad.lean | 2 -- Mathlib/Data/Set/Finite/Range.lean | 3 +- Mathlib/Dynamics/PeriodicPts.lean | 6 ++-- .../NumberField/Discriminant/Basic.lean | 8 ++++- .../NumberField/Units/DirichletTheorem.lean | 19 +++++++---- .../RankAndCardinality.lean | 2 -- .../TranscendenceBasis.lean | 2 -- .../AlgebraicIndependent/Transcendental.lean | 2 -- .../DedekindDomain/AdicValuation.lean | 12 +++++-- .../RingTheory/DedekindDomain/Different.lean | 10 +++--- .../DedekindDomain/Factorization.lean | 2 +- Mathlib/RingTheory/DedekindDomain/Ideal.lean | 3 +- Mathlib/RingTheory/Flat/FaithfullyFlat.lean | 7 ++-- Mathlib/RingTheory/HahnSeries/Addition.lean | 7 ++-- .../IsIntegralClosure/Basic.lean | 1 - Mathlib/RingTheory/LocalProperties/Basic.lean | 2 +- Mathlib/RingTheory/Localization/Integral.lean | 4 +-- .../Polynomial/IrreducibleRing.lean | 2 -- .../Polynomial/SeparableDegree.lean | 1 - Mathlib/RingTheory/PowerSeries/Inverse.lean | 6 ++-- Mathlib/RingTheory/RingHom/Finite.lean | 4 ++- Mathlib/RingTheory/RingHom/FiniteType.lean | 6 +++- Mathlib/RingTheory/RootsOfUnity/Basic.lean | 10 +++--- .../RootsOfUnity/PrimitiveRoots.lean | 9 ++++-- Mathlib/RingTheory/Smooth/StandardSmooth.lean | 10 +++++- .../NormalizedFactors.lean | 3 +- .../Valuation/ValuationSubring.lean | 8 ++--- Mathlib/RingTheory/WittVector/Domain.lean | 3 +- Mathlib/RingTheory/WittVector/InitTail.lean | 4 +-- Mathlib/Topology/EMetricSpace/Diam.lean | 2 +- 113 files changed, 339 insertions(+), 307 deletions(-) diff --git a/Mathlib/Analysis/Analytic/CPolynomial.lean b/Mathlib/Analysis/Analytic/CPolynomial.lean index 0fe03f76749d9..8c2c58d2957d5 100644 --- a/Mathlib/Analysis/Analytic/CPolynomial.lean +++ b/Mathlib/Analysis/Analytic/CPolynomial.lean @@ -47,7 +47,7 @@ analytic. variable {𝕜 E F G : Type*} [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace 𝕜 F] [NormedAddCommGroup G] [NormedSpace 𝕜 G] -open scoped Classical Topology +open scoped Topology open Set Filter Asymptotics NNReal ENNReal variable {f g : E → F} {p pf pg : FormalMultilinearSeries 𝕜 E F} {x : E} {r r' : ℝ≥0∞} {n m : ℕ} diff --git a/Mathlib/Analysis/Analytic/Constructions.lean b/Mathlib/Analysis/Analytic/Constructions.lean index 3fcad68970d2c..a1a029bcd24a3 100644 --- a/Mathlib/Analysis/Analytic/Constructions.lean +++ b/Mathlib/Analysis/Analytic/Constructions.lean @@ -21,7 +21,7 @@ We show that the following are analytic: noncomputable section -open scoped Classical Topology +open scoped Topology open Filter Asymptotics ENNReal NNReal variable {α : Type*} @@ -883,6 +883,7 @@ theorem AnalyticOnNhd.div {f g : E → 𝕝} {s : Set E} theorem Finset.analyticWithinAt_sum {f : α → E → F} {c : E} {s : Set E} (N : Finset α) (h : ∀ n ∈ N, AnalyticWithinAt 𝕜 (f n) s c) : AnalyticWithinAt 𝕜 (fun z ↦ ∑ n ∈ N, f n z) s c := by + classical induction' N using Finset.induction with a B aB hB · simp only [Finset.sum_empty] exact analyticWithinAt_const @@ -916,6 +917,7 @@ theorem Finset.analyticOnNhd_sum {f : α → E → F} {s : Set E} theorem Finset.analyticWithinAt_prod {A : Type*} [NormedCommRing A] [NormedAlgebra 𝕜 A] {f : α → E → A} {c : E} {s : Set E} (N : Finset α) (h : ∀ n ∈ N, AnalyticWithinAt 𝕜 (f n) s c) : AnalyticWithinAt 𝕜 (fun z ↦ ∏ n ∈ N, f n z) s c := by + classical induction' N using Finset.induction with a B aB hB · simp only [Finset.prod_empty] exact analyticWithinAt_const diff --git a/Mathlib/Analysis/Analytic/IsolatedZeros.lean b/Mathlib/Analysis/Analytic/IsolatedZeros.lean index 6cb55568fc2c7..7b5636ce8c665 100644 --- a/Mathlib/Analysis/Analytic/IsolatedZeros.lean +++ b/Mathlib/Analysis/Analytic/IsolatedZeros.lean @@ -25,9 +25,6 @@ useful in this setup. accumulation point in `U` then `f` is identically `0` on `U`. -/ - -open scoped Classical - open Filter Function Nat FormalMultilinearSeries EMetric Set open scoped Topology @@ -192,6 +189,7 @@ theorem exists_eventuallyEq_pow_smul_nonzero_iff (hf : AnalyticAt 𝕜 f z₀) : hp.iterate_dslope_fslope_ne_zero (hf_ne.imp hp.locally_zero_iff.mpr), hp.eq_pow_order_mul_iterate_dslope⟩ +open scoped Classical in /-- The order of vanishing of `f` at `z₀`, as an element of `ℕ∞`. This is defined to be `∞` if `f` is identically 0 on a neighbourhood of `z₀`, and otherwise the diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index c9ef3caef6893..74755cbeef3a1 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -45,7 +45,6 @@ it suffices to assume that `f` is zero wherever `g` is. (This generalization is the Fréchet derivative.) -/ -open scoped Classical open Set Topology Filter NNReal namespace Asymptotics @@ -1571,6 +1570,7 @@ variable {ι : Type*} {A : ι → α → E'} {C : ι → ℝ} {s : Finset ι} theorem IsBigOWith.sum (h : ∀ i ∈ s, IsBigOWith (C i) l (A i) g) : IsBigOWith (∑ i ∈ s, C i) l (fun x => ∑ i ∈ s, A i x) g := by + classical induction' s using Finset.induction_on with i s is IH · simp only [isBigOWith_zero', Finset.sum_empty, forall_true_iff] · simp only [is, Finset.sum_insert, not_false_iff] @@ -1582,6 +1582,7 @@ theorem IsBigO.sum (h : ∀ i ∈ s, A i =O[l] g) : (fun x => ∑ i ∈ s, A i x exact ⟨_, IsBigOWith.sum hC⟩ theorem IsLittleO.sum (h : ∀ i ∈ s, A i =o[l] g') : (fun x => ∑ i ∈ s, A i x) =o[l] g' := by + classical induction' s using Finset.induction_on with i s is IH · simp only [isLittleO_zero, Finset.sum_empty, forall_true_iff] · simp only [is, Finset.sum_insert, not_false_iff] @@ -1855,8 +1856,9 @@ theorem bound_of_isBigO_cofinite (h : f =O[cofinite] g'') : exact fun hx => (div_le_iff₀ (norm_pos_iff.2 h₀)).1 (this _ hx) theorem isBigO_cofinite_iff (h : ∀ x, g'' x = 0 → f'' x = 0) : - f'' =O[cofinite] g'' ↔ ∃ C, ∀ x, ‖f'' x‖ ≤ C * ‖g'' x‖ := - ⟨fun h' => + f'' =O[cofinite] g'' ↔ ∃ C, ∀ x, ‖f'' x‖ ≤ C * ‖g'' x‖ := by + classical + exact ⟨fun h' => let ⟨C, _C₀, hC⟩ := bound_of_isBigO_cofinite h' ⟨C, fun x => if hx : g'' x = 0 then by simp [h _ hx, hx] else hC hx⟩, fun h => (isBigO_top.2 h).mono le_top⟩ diff --git a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean index 622d8b9eb3b5d..1700dfa697aab 100644 --- a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean @@ -55,7 +55,7 @@ open Set Function Metric Filter noncomputable section -open scoped Classical NNReal Topology +open scoped NNReal Topology namespace BoxIntegral @@ -273,6 +273,7 @@ theorem withBotCoe_subset_iff {I J : WithBot (Box ι)} : (I : Set (ι → ℝ)) theorem withBotCoe_inj {I J : WithBot (Box ι)} : (I : Set (ι → ℝ)) = J ↔ I = J := by simp only [Subset.antisymm_iff, ← le_antisymm_iff, withBotCoe_subset_iff] +open scoped Classical in /-- Make a `WithBot (Box ι)` from a pair of corners `l u : ι → ℝ`. If `l i < u i` for all `i`, then the result is `⟨l, u, _⟩ : Box ι`, otherwise it is `⊥`. In any case, the result interpreted as a set in `ι → ℝ` is the set `{x : ι → ℝ | ∀ i, x i ∈ Ioc (l i) (u i)}`. -/ diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean b/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean index adb77e996ca75..3c04cb6f9c6d6 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Additive.lean @@ -28,7 +28,6 @@ rectangular box, additive function noncomputable section -open scoped Classical open Function Set namespace BoxIntegral @@ -115,6 +114,7 @@ def ofMapSplitAdd [Finite ι] (f : Box ι → M) (I₀ : WithTop (Box ι)) (hf : ∀ I : Box ι, ↑I ≤ I₀ → ∀ {i x}, x ∈ Ioo (I.lower i) (I.upper i) → (I.splitLower i x).elim' 0 f + (I.splitUpper i x).elim' 0 f = f I) : ι →ᵇᵃ[I₀] M := by + classical refine ⟨f, ?_⟩ replace hf : ∀ I : Box ι, ↑I ≤ I₀ → ∀ s, (∑ J ∈ (splitMany I s).boxes, f J) = f I := by intro I hI s diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean b/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean index a381630fbe12e..03d14e71d483f 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean @@ -37,7 +37,7 @@ rectangular box, partition -/ open Set Finset Function -open scoped Classical NNReal +open scoped NNReal noncomputable section @@ -174,6 +174,7 @@ theorem injOn_setOf_mem_Icc_setOf_lower_eq (x : ι → ℝ) : · have hi₂ : J₂.lower i < x i := (hx₂.1 i).lt_of_ne (mt (H _).2 hi₁.ne) exact ⟨x i, ⟨hi₁, hx₁.2 i⟩, ⟨hi₂, hx₂.2 i⟩⟩ +open scoped Classical in /-- The set of boxes of a prepartition that contain `x` in their closures has cardinality at most `2 ^ Fintype.card ι`. -/ theorem card_filter_mem_Icc_le [Fintype ι] (x : ι → ℝ) : @@ -249,6 +250,7 @@ theorem eq_of_boxes_subset_iUnion_superset (h₁ : π₁.boxes ⊆ π₂.boxes) ⟨fun _ hJ₁ _ hJ₂ Hne => (π₂.eq_of_mem_of_mem hJ₁ (h₁ hJ₂) Hne.choose_spec.1 Hne.choose_spec.2).le, h₂⟩ +open scoped Classical in /-- Given a prepartition `π` of a box `I` and a collection of prepartitions `πi J` of all boxes `J ∈ π`, returns the prepartition of `I` into the union of the boxes of all `πi J`. @@ -300,6 +302,7 @@ theorem biUnion_congr_of_le (h : π₁ = π₂) (hi : ∀ J ≤ I, πi₁ J = π theorem iUnion_biUnion (πi : ∀ J : Box ι, Prepartition J) : (π.biUnion πi).iUnion = ⋃ J ∈ π, (πi J).iUnion := by simp [Prepartition.iUnion] +open scoped Classical in @[simp] theorem sum_biUnion_boxes {M : Type*} [AddCommMonoid M] (π : Prepartition I) (πi : ∀ J, Prepartition J) (f : Box ι → M) : @@ -308,6 +311,7 @@ theorem sum_biUnion_boxes {M : Type*} [AddCommMonoid M] (π : Prepartition I) refine Finset.sum_biUnion fun J₁ h₁ J₂ h₂ hne => Finset.disjoint_left.2 fun J' h₁' h₂' => ?_ exact hne (π.eq_of_le_of_le h₁ h₂ ((πi J₁).le_of_mem h₁') ((πi J₂).le_of_mem h₂')) +open scoped Classical in /-- Given a box `J ∈ π.biUnion πi`, returns the box `J' ∈ π` such that `J ∈ πi J'`. For `J ∉ π.biUnion πi`, returns `I`. -/ def biUnionIndex (πi : ∀ (J : Box ι), Prepartition J) (J : Box ι) : Box ι := @@ -410,6 +414,7 @@ theorem sum_ofWithBot {M : Type*} [AddCommMonoid M] (boxes : Finset (WithBot (Bo ∑ J ∈ boxes, Option.elim' 0 f J := Finset.sum_eraseNone _ _ +open scoped Classical in /-- Restrict a prepartition to a box. -/ def restrict (π : Prepartition I) (J : Box ι) : Prepartition J := ofWithBot (π.boxes.image fun J' : Box ι => J ⊓ J') @@ -433,6 +438,7 @@ theorem mem_restrict' : J₁ ∈ π.restrict J ↔ ∃ J' ∈ π, (J₁ : Set ( @[mono] theorem restrict_mono {π₁ π₂ : Prepartition I} (Hle : π₁ ≤ π₂) : π₁.restrict J ≤ π₂.restrict J := by + classical refine ofWithBot_mono fun J₁ hJ₁ hne => ?_ rw [Finset.mem_image] at hJ₁; rcases hJ₁ with ⟨J₁, hJ₁, rfl⟩ rcases Hle hJ₁ with ⟨J₂, hJ₂, hle⟩ @@ -444,6 +450,7 @@ theorem monotone_restrict : Monotone fun π : Prepartition I => restrict π J := /-- Restricting to a larger box does not change the set of boxes. We cannot claim equality of prepartitions because they have different types. -/ theorem restrict_boxes_of_le (π : Prepartition I) (h : I ≤ J) : (π.restrict J).boxes = π.boxes := by + classical simp only [restrict, ofWithBot, eraseNone_eq_biUnion] refine Finset.image_biUnion.trans ?_ refine (Finset.biUnion_congr rfl ?_).trans Finset.biUnion_singleton_eq_self @@ -510,6 +517,7 @@ theorem mem_inf {π₁ π₂ : Prepartition I} : theorem iUnion_inf (π₁ π₂ : Prepartition I) : (π₁ ⊓ π₂).iUnion = π₁.iUnion ∩ π₂.iUnion := by simp only [inf_def, iUnion_biUnion, iUnion_restrict, ← iUnion_inter, ← iUnion_def] +open scoped Classical in /-- The prepartition with boxes `{J ∈ π | p J}`. -/ @[simps] def filter (π : Prepartition I) (p : Box ι → Prop) : Prepartition I where @@ -518,8 +526,9 @@ def filter (π : Prepartition I) (p : Box ι → Prop) : Prepartition I where pairwiseDisjoint _ h₁ _ h₂ := π.disjoint_coe_of_mem (mem_filter.1 h₁).1 (mem_filter.1 h₂).1 @[simp] -theorem mem_filter {p : Box ι → Prop} : J ∈ π.filter p ↔ J ∈ π ∧ p J := - Finset.mem_filter +theorem mem_filter {p : Box ι → Prop} : J ∈ π.filter p ↔ J ∈ π ∧ p J := by + classical + exact Finset.mem_filter theorem filter_le (π : Prepartition I) (p : Box ι → Prop) : π.filter p ≤ π := fun J hJ => let ⟨hπ, _⟩ := π.mem_filter.1 hJ @@ -545,11 +554,13 @@ theorem iUnion_filter_not (π : Prepartition I) (p : Box ι → Prop) : rw [Set.union_eq_left, filter_boxes, coe_filter] exact fun _ ⟨h, _⟩ => h +open scoped Classical in theorem sum_fiberwise {α M} [AddCommMonoid M] (π : Prepartition I) (f : Box ι → α) (g : Box ι → M) : (∑ y ∈ π.boxes.image f, ∑ J ∈ (π.filter fun J => f J = y).boxes, g J) = ∑ J ∈ π.boxes, g J := by convert sum_fiberwise_of_maps_to (fun _ => Finset.mem_image_of_mem f) g +open scoped Classical in /-- Union of two disjoint prepartitions. -/ @[simps] def disjUnion (π₁ π₂ : Prepartition I) (h : Disjoint π₁.iUnion π₂.iUnion) : Prepartition I where @@ -562,14 +573,15 @@ def disjUnion (π₁ π₂ : Prepartition I) (h : Disjoint π₁.iUnion π₂.iU @[simp] theorem mem_disjUnion (H : Disjoint π₁.iUnion π₂.iUnion) : - J ∈ π₁.disjUnion π₂ H ↔ J ∈ π₁ ∨ J ∈ π₂ := - Finset.mem_union + J ∈ π₁.disjUnion π₂ H ↔ J ∈ π₁ ∨ J ∈ π₂ := by + classical exact Finset.mem_union @[simp] theorem iUnion_disjUnion (h : Disjoint π₁.iUnion π₂.iUnion) : (π₁.disjUnion π₂ h).iUnion = π₁.iUnion ∪ π₂.iUnion := by simp [disjUnion, Prepartition.iUnion, iUnion_or, iUnion_union_distrib] +open scoped Classical in @[simp] theorem sum_disj_union_boxes {M : Type*} [AddCommMonoid M] (h : Disjoint π₁.iUnion π₂.iUnion) (f : Box ι → M) : @@ -592,13 +604,13 @@ theorem distortion_le_iff {c : ℝ≥0} : π.distortion ≤ c ↔ ∀ J ∈ π, Finset.sup_le_iff theorem distortion_biUnion (π : Prepartition I) (πi : ∀ J, Prepartition J) : - (π.biUnion πi).distortion = π.boxes.sup fun J => (πi J).distortion := - sup_biUnion _ _ + (π.biUnion πi).distortion = π.boxes.sup fun J => (πi J).distortion := by + classical exact sup_biUnion _ _ @[simp] theorem distortion_disjUnion (h : Disjoint π₁.iUnion π₂.iUnion) : - (π₁.disjUnion π₂ h).distortion = max π₁.distortion π₂.distortion := - sup_union + (π₁.disjUnion π₂ h).distortion = max π₁.distortion π₂.distortion := by + classical exact sup_union theorem distortion_of_const {c} (h₁ : π.boxes.Nonempty) (h₂ : ∀ J ∈ π, Box.distortion J = c) : π.distortion = c := diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean index f9ca9c7272245..d21e854c59b7a 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean @@ -164,7 +164,7 @@ integral, rectangular box, partition, filter -/ open Set Function Filter Metric Finset Bool -open scoped Classical Topology Filter NNReal +open scoped Topology Filter NNReal noncomputable section @@ -366,6 +366,7 @@ 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 + classical 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, fun hD => (distortion_filter_le _ _).trans (hπ.3 hD), fun hD => ?_⟩ rcases hπ.4 hD with ⟨π₁, hπ₁U, hc⟩ diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Split.lean b/Mathlib/Analysis/BoxIntegral/Partition/Split.lean index 3e39d436eeb97..99d1ca38f0e7c 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Split.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Split.lean @@ -37,7 +37,6 @@ rectangular box, partition, hyperplane noncomputable section -open scoped Classical open Function Set Filter namespace BoxIntegral @@ -48,6 +47,7 @@ namespace Box variable {I : Box ι} {i : ι} {x : ℝ} {y : ι → ℝ} +open scoped Classical in /-- Given a box `I` and `x ∈ (I.lower i, I.upper i)`, the hyperplane `{y : ι → ℝ | y i = x}` splits `I` into two boxes. `BoxIntegral.Box.splitLower I i x` is the box `I ∩ {y | y i ≤ x}` (if it is nonempty). As usual, we represent a box that may be empty as @@ -68,6 +68,7 @@ theorem splitLower_le : I.splitLower i x ≤ I := @[simp] theorem splitLower_eq_bot {i x} : I.splitLower i x = ⊥ ↔ x ≤ I.lower i := by + classical rw [splitLower, mk'_eq_bot, exists_update_iff I.upper fun j y => y ≤ I.lower j] simp [(I.lower_lt_upper _).not_le] @@ -83,6 +84,7 @@ theorem splitLower_def [DecidableEq ι] {i x} (h : x ∈ Ioo (I.lower i) (I.uppe simp (config := { unfoldPartialApp := true }) only [splitLower, mk'_eq_coe, min_eq_left h.2.le, update, and_self] +open scoped Classical in /-- Given a box `I` and `x ∈ (I.lower i, I.upper i)`, the hyperplane `{y : ι → ℝ | y i = x}` splits `I` into two boxes. `BoxIntegral.Box.splitUpper I i x` is the box `I ∩ {y | x < y i}` (if it is nonempty). As usual, we represent a box that may be empty as @@ -92,6 +94,7 @@ def splitUpper (I : Box ι) (i : ι) (x : ℝ) : WithBot (Box ι) := @[simp] theorem coe_splitUpper : (splitUpper I i x : Set (ι → ℝ)) = ↑I ∩ { y | x < y i } := by + classical rw [splitUpper, coe_mk'] ext y simp only [mem_univ_pi, mem_Ioc, mem_inter_iff, mem_coe, mem_setOf_eq, forall_and, @@ -104,6 +107,7 @@ theorem splitUpper_le : I.splitUpper i x ≤ I := @[simp] theorem splitUpper_eq_bot {i x} : I.splitUpper i x = ⊥ ↔ I.upper i ≤ x := by + classical rw [splitUpper, mk'_eq_bot, exists_update_iff I.lower fun j y => I.upper j ≤ y] simp [(I.lower_lt_upper _).not_le] @@ -140,6 +144,7 @@ namespace Prepartition variable {I J : Box ι} {i : ι} {x : ℝ} +open scoped Classical in /-- The partition of `I : Box ι` into the boxes `I ∩ {y | y ≤ x i}` and `I ∩ {y | x i < y}`. One of these boxes can be empty, then this partition is just the single-box partition `⊤`. -/ def split (I : Box ι) (i : ι) (x : ℝ) : Prepartition I := @@ -173,6 +178,7 @@ theorem isPartitionSplit (I : Box ι) (i : ι) (x : ℝ) : IsPartition (split I theorem sum_split_boxes {M : Type*} [AddCommMonoid M] (I : Box ι) (i : ι) (x : ℝ) (f : Box ι → M) : (∑ J ∈ (split I i x).boxes, f J) = (I.splitLower i x).elim' 0 f + (I.splitUpper i x).elim' 0 f := by + classical rw [split, sum_ofWithBot, Finset.sum_pair (I.splitLower_ne_splitUpper i x)] /-- If `x ∉ (I.lower i, I.upper i)`, then the hyperplane `{y | y i = x}` does not split `I`. -/ @@ -218,6 +224,7 @@ def splitMany (I : Box ι) (s : Finset (ι × ℝ)) : Prepartition I := theorem splitMany_empty (I : Box ι) : splitMany I ∅ = ⊤ := Finset.inf_empty +open scoped Classical in @[simp] theorem splitMany_insert (I : Box ι) (s : Finset (ι × ℝ)) (p : ι × ℝ) : splitMany I (insert p s) = splitMany I s ⊓ split I p.1 p.2 := by @@ -227,8 +234,9 @@ theorem splitMany_le_split (I : Box ι) {s : Finset (ι × ℝ)} {p : ι × ℝ} splitMany I s ≤ split I p.1 p.2 := Finset.inf_le hp -theorem isPartition_splitMany (I : Box ι) (s : Finset (ι × ℝ)) : IsPartition (splitMany I s) := - Finset.induction_on s (by simp only [splitMany_empty, isPartitionTop]) fun a s _ hs => by +theorem isPartition_splitMany (I : Box ι) (s : Finset (ι × ℝ)) : IsPartition (splitMany I s) := by + classical + exact Finset.induction_on s (by simp only [splitMany_empty, isPartitionTop]) fun a s _ hs => by simpa only [splitMany_insert, inf_split] using hs.biUnion fun J _ => isPartitionSplit _ _ _ @[simp] @@ -237,10 +245,12 @@ theorem iUnion_splitMany (I : Box ι) (s : Finset (ι × ℝ)) : (splitMany I s) theorem inf_splitMany {I : Box ι} (π : Prepartition I) (s : Finset (ι × ℝ)) : π ⊓ splitMany I s = π.biUnion fun J => splitMany J s := by + classical induction' s using Finset.induction_on with p s _ ihp · simp · simp_rw [splitMany_insert, ← inf_assoc, ihp, inf_split, biUnion_assoc] +open scoped Classical in /-- Let `s : Finset (ι × ℝ)` be a set of hyperplanes `{x : ι → ℝ | x i = r}` in `ι → ℝ` encoded as pairs `(i, r)`. Suppose that this set contains all faces of a box `J`. The hyperplanes of `s` split a box `I` into subboxes. Let `Js` be one of them. If `J` and `Js` have nonempty intersection, then @@ -272,6 +282,7 @@ intersection, then `J' ≤ J`. -/ theorem eventually_not_disjoint_imp_le_of_mem_splitMany (s : Finset (Box ι)) : ∀ᶠ t : Finset (ι × ℝ) in atTop, ∀ (I : Box ι), ∀ J ∈ s, ∀ J' ∈ splitMany I t, ¬Disjoint (J : WithBot (Box ι)) J' → J' ≤ J := by + classical cases nonempty_fintype ι refine eventually_atTop.2 ⟨s.biUnion fun J => Finset.univ.biUnion fun i => {(i, J.lower i), (i, J.upper i)}, diff --git a/Mathlib/Analysis/BoxIntegral/Partition/SubboxInduction.lean b/Mathlib/Analysis/BoxIntegral/Partition/SubboxInduction.lean index 2e7c804620e0d..58f1cc837f67c 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/SubboxInduction.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/SubboxInduction.lean @@ -31,7 +31,6 @@ namespace BoxIntegral open Set Metric -open scoped Classical open Topology noncomputable section @@ -212,6 +211,7 @@ theorem isPartition_unionComplToSubordinate (π₁ : TaggedPrepartition I) (π IsPartition (π₁.unionComplToSubordinate π₂ hU r) := Prepartition.isPartitionDisjUnionOfEqDiff ((π₂.iUnion_toSubordinate r).trans hU) +open scoped Classical in @[simp] theorem unionComplToSubordinate_boxes (π₁ : TaggedPrepartition I) (π₂ : Prepartition I) (hU : π₂.iUnion = ↑I \ π₁.iUnion) (r : (ι → ℝ) → Ioi (0 : ℝ)) : diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Tagged.lean b/Mathlib/Analysis/BoxIntegral/Partition/Tagged.lean index 9710dd56108b3..5b3a16866e31c 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Tagged.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Tagged.lean @@ -26,7 +26,6 @@ rectangular box, box partition noncomputable section open Finset Function ENNReal NNReal Set -open scoped Classical namespace BoxIntegral @@ -92,8 +91,8 @@ def filter (p : Box ι → Prop) : TaggedPrepartition I := ⟨π.1.filter p, π.2, π.3⟩ @[simp] -theorem mem_filter {p : Box ι → Prop} : J ∈ π.filter p ↔ J ∈ π ∧ p J := - Finset.mem_filter +theorem mem_filter {p : Box ι → Prop} : J ∈ π.filter p ↔ J ∈ π ∧ p J := by + classical exact Finset.mem_filter @[simp] theorem iUnion_filter_not (π : TaggedPrepartition I) (p : Box ι → Prop) : @@ -202,7 +201,8 @@ theorem isHenstock_biUnionTagged {π : Prepartition I} {πi : ∀ J, TaggedPrepa /-- In a Henstock prepartition, there are at most `2 ^ Fintype.card ι` boxes with a given tag. -/ theorem IsHenstock.card_filter_tag_eq_le [Fintype ι] (h : π.IsHenstock) (x : ι → ℝ) : - #{J ∈ π.boxes | π.tag J = x} ≤ 2 ^ Fintype.card ι := + #{J ∈ π.boxes | π.tag J = x} ≤ 2 ^ Fintype.card ι := by + classical calc #{J ∈ π.boxes | π.tag J = x} ≤ #{J ∈ π.boxes | x ∈ Box.Icc J} := by refine Finset.card_le_card fun J hJ => ?_ @@ -287,6 +287,7 @@ theorem isSubordinate_single [Fintype ι] (hJ : J ≤ I) (h : x ∈ Box.Icc I) : theorem iUnion_single (hJ : J ≤ I) (h : x ∈ Box.Icc I) : (single I J hJ x h).iUnion = J := Prepartition.iUnion_single hJ +open scoped Classical in /-- Union of two tagged prepartitions with disjoint unions of boxes. -/ def disjUnion (π₁ π₂ : TaggedPrepartition I) (h : Disjoint π₁.iUnion π₂.iUnion) : TaggedPrepartition I where @@ -297,14 +298,15 @@ def disjUnion (π₁ π₂ : TaggedPrepartition I) (h : Disjoint π₁.iUnion π split_ifs exacts [π₁.tag_mem_Icc J, π₂.tag_mem_Icc J] +open scoped Classical in @[simp] theorem disjUnion_boxes (h : Disjoint π₁.iUnion π₂.iUnion) : (π₁.disjUnion π₂ h).boxes = π₁.boxes ∪ π₂.boxes := rfl @[simp] theorem mem_disjUnion (h : Disjoint π₁.iUnion π₂.iUnion) : - J ∈ π₁.disjUnion π₂ h ↔ J ∈ π₁ ∨ J ∈ π₂ := - Finset.mem_union + J ∈ π₁.disjUnion π₂ h ↔ J ∈ π₁ ∨ J ∈ π₂ := by + classical exact Finset.mem_union @[simp] theorem iUnion_disjUnion (h : Disjoint π₁.iUnion π₂.iUnion) : @@ -321,6 +323,7 @@ theorem disjUnion_tag_of_mem_right (h : Disjoint π₁.iUnion π₂.iUnion) (hJ theorem IsSubordinate.disjUnion [Fintype ι] (h₁ : IsSubordinate π₁ r) (h₂ : IsSubordinate π₂ r) (h : Disjoint π₁.iUnion π₂.iUnion) : IsSubordinate (π₁.disjUnion π₂ h) r := by + classical refine fun J hJ => (Finset.mem_union.1 hJ).elim (fun hJ => ?_) fun hJ => ?_ · rw [disjUnion_tag_of_mem_left _ hJ] exact h₁ _ hJ @@ -329,6 +332,7 @@ theorem IsSubordinate.disjUnion [Fintype ι] (h₁ : IsSubordinate π₁ r) (h theorem IsHenstock.disjUnion (h₁ : IsHenstock π₁) (h₂ : IsHenstock π₂) (h : Disjoint π₁.iUnion π₂.iUnion) : IsHenstock (π₁.disjUnion π₂ h) := by + classical refine fun J hJ => (Finset.mem_union.1 hJ).elim (fun hJ => ?_) fun hJ => ?_ · rw [disjUnion_tag_of_mem_left _ hJ] exact h₁ _ hJ @@ -364,18 +368,18 @@ theorem distortion_le_iff {c : ℝ≥0} : π.distortion ≤ c ↔ ∀ J ∈ π, @[simp] theorem _root_.BoxIntegral.Prepartition.distortion_biUnionTagged (π : Prepartition I) (πi : ∀ J, TaggedPrepartition J) : - (π.biUnionTagged πi).distortion = π.boxes.sup fun J => (πi J).distortion := - sup_biUnion _ _ + (π.biUnionTagged πi).distortion = π.boxes.sup fun J => (πi J).distortion := by + classical exact sup_biUnion _ _ @[simp] theorem distortion_biUnionPrepartition (π : TaggedPrepartition I) (πi : ∀ J, Prepartition J) : - (π.biUnionPrepartition πi).distortion = π.boxes.sup fun J => (πi J).distortion := - sup_biUnion _ _ + (π.biUnionPrepartition πi).distortion = π.boxes.sup fun J => (πi J).distortion := by + classical exact sup_biUnion _ _ @[simp] theorem distortion_disjUnion (h : Disjoint π₁.iUnion π₂.iUnion) : - (π₁.disjUnion π₂ h).distortion = max π₁.distortion π₂.distortion := - sup_union + (π₁.disjUnion π₂ h).distortion = max π₁.distortion π₂.distortion := by + classical exact sup_union theorem distortion_of_const {c} (h₁ : π.boxes.Nonempty) (h₂ : ∀ J ∈ π, Box.distortion J = c) : π.distortion = c := @@ -386,8 +390,8 @@ theorem distortion_single (hJ : J ≤ I) (h : x ∈ Box.Icc I) : distortion (single I J hJ x h) = J.distortion := sup_singleton -theorem distortion_filter_le (p : Box ι → Prop) : (π.filter p).distortion ≤ π.distortion := - sup_mono (filter_subset _ _) +theorem distortion_filter_le (p : Box ι → Prop) : (π.filter p).distortion ≤ π.distortion := by + classical exact sup_mono (filter_subset _ _) end Distortion diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index e772546295c8f..6fbd690d334d6 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -96,7 +96,6 @@ derivative, differentiability, higher derivative, `C^n`, multilinear, Taylor ser noncomputable section -open scoped Classical open NNReal Topology Filter /- diff --git a/Mathlib/Analysis/Calculus/Darboux.lean b/Mathlib/Analysis/Calculus/Darboux.lean index dfd41fabe3597..a439370033e45 100644 --- a/Mathlib/Analysis/Calculus/Darboux.lean +++ b/Mathlib/Analysis/Calculus/Darboux.lean @@ -17,7 +17,7 @@ intermediate values. The proof is based on the open Filter Set -open scoped Topology Classical +open scoped Topology variable {a b : ℝ} {f f' : ℝ → ℝ} diff --git a/Mathlib/Analysis/Calculus/Deriv/Add.lean b/Mathlib/Analysis/Calculus/Deriv/Add.lean index b5987de05ae42..3d6b994ca666b 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Add.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Add.lean @@ -22,7 +22,6 @@ derivative universe u v w -open scoped Classical open scoped Topology Filter ENNReal open Asymptotics Set diff --git a/Mathlib/Analysis/Calculus/Deriv/Comp.lean b/Mathlib/Analysis/Calculus/Deriv/Comp.lean index 979caef8106a3..015bd0a24acb8 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Comp.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Comp.lean @@ -34,7 +34,7 @@ derivative, chain rule universe u v w -open scoped Classical Topology Filter ENNReal +open scoped Topology Filter ENNReal open Filter Asymptotics Set diff --git a/Mathlib/Analysis/Calculus/Deriv/Prod.lean b/Mathlib/Analysis/Calculus/Deriv/Prod.lean index dfd2e896128ad..a6a1adbcdb78d 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Prod.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Prod.lean @@ -22,10 +22,7 @@ derivative universe u v w -open scoped Classical -open Topology Filter - -open Filter Asymptotics Set +open Topology Filter Asymptotics Set variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] variable {F : Type v} [NormedAddCommGroup F] [NormedSpace 𝕜 F] diff --git a/Mathlib/Analysis/Calculus/Deriv/ZPow.lean b/Mathlib/Analysis/Calculus/Deriv/ZPow.lean index 73cd119155a71..065ffb0ec882e 100644 --- a/Mathlib/Analysis/Calculus/Deriv/ZPow.lean +++ b/Mathlib/Analysis/Calculus/Deriv/ZPow.lean @@ -21,10 +21,7 @@ derivative, power universe u v w -open scoped Classical -open Topology Filter - -open Filter Asymptotics Set +open Topology Filter Asymptotics Set variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] variable {E : Type v} [NormedAddCommGroup E] [NormedSpace 𝕜 E] diff --git a/Mathlib/Analysis/Calculus/FDeriv/Add.lean b/Mathlib/Analysis/Calculus/FDeriv/Add.lean index fd877bcc04666..2677296a22c2e 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Add.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Add.lean @@ -23,8 +23,6 @@ This file contains the usual formulas (and existence assertions) for the derivat open Filter Asymptotics ContinuousLinearMap -open scoped Classical - noncomputable section section @@ -229,10 +227,11 @@ theorem differentiable_add_const_iff (c : F) : ⟨fun h => by simpa using h.add_const (-c), fun h => h.add_const c⟩ theorem fderivWithin_add_const (hxs : UniqueDiffWithinAt 𝕜 s x) (c : F) : - fderivWithin 𝕜 (fun y => f y + c) s x = fderivWithin 𝕜 f s x := - if hf : DifferentiableWithinAt 𝕜 f s x then (hf.hasFDerivWithinAt.add_const c).fderivWithin hxs - else by - rw [fderivWithin_zero_of_not_differentiableWithinAt hf, + fderivWithin 𝕜 (fun y => f y + c) s x = fderivWithin 𝕜 f s x := by + classical + by_cases hf : DifferentiableWithinAt 𝕜 f s x + · exact (hf.hasFDerivWithinAt.add_const c).fderivWithin hxs + · rw [fderivWithin_zero_of_not_differentiableWithinAt hf, fderivWithin_zero_of_not_differentiableWithinAt] simpa @@ -424,10 +423,11 @@ theorem differentiable_neg_iff : (Differentiable 𝕜 fun y => -f y) ↔ Differe ⟨fun h => by simpa only [neg_neg] using h.neg, fun h => h.neg⟩ theorem fderivWithin_neg (hxs : UniqueDiffWithinAt 𝕜 s x) : - fderivWithin 𝕜 (fun y => -f y) s x = -fderivWithin 𝕜 f s x := - if h : DifferentiableWithinAt 𝕜 f s x then h.hasFDerivWithinAt.neg.fderivWithin hxs - else by - rw [fderivWithin_zero_of_not_differentiableWithinAt h, + fderivWithin 𝕜 (fun y => -f y) s x = -fderivWithin 𝕜 f s x := by + classical + by_cases h : DifferentiableWithinAt 𝕜 f s x + · exact h.hasFDerivWithinAt.neg.fderivWithin hxs + · rw [fderivWithin_zero_of_not_differentiableWithinAt h, fderivWithin_zero_of_not_differentiableWithinAt, neg_zero] simpa @@ -739,6 +739,7 @@ theorem differentiableWithinAt_comp_add_right (a : E) : theorem fderivWithin_comp_add_right (a : E) : fderivWithin 𝕜 (fun x ↦ f (x + a)) s x = fderivWithin 𝕜 f (a +ᵥ s) (x + a) := by + classical simp only [fderivWithin, hasFDerivWithinAt_comp_add_right, DifferentiableWithinAt] theorem hasFDerivWithinAt_comp_add_left (a : E) : diff --git a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean index ee68296873c8a..be3bc501cfa89 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean @@ -120,10 +120,7 @@ derivative, differentiable, Fréchet, calculus -/ -open Filter Asymptotics ContinuousLinearMap Set Metric - -open scoped Classical -open Topology NNReal Filter Asymptotics ENNReal +open Filter Asymptotics ContinuousLinearMap Set Metric Topology NNReal ENNReal noncomputable section @@ -178,6 +175,7 @@ non-unique). -/ def DifferentiableAt (f : E → F) (x : E) := ∃ f' : E →L[𝕜] F, HasFDerivAt f f' x +open scoped Classical in /-- If `f` has a derivative at `x` within `s`, then `fderivWithin 𝕜 f s x` is such a derivative. Otherwise, it is set to `0`. We also set it to be zero, if zero is one of possible derivatives. -/ irreducible_def fderivWithin (f : E → F) (s : Set E) (x : E) : E →L[𝕜] F := @@ -623,6 +621,7 @@ theorem fderivWithin_subset (st : s ⊆ t) (ht : UniqueDiffWithinAt 𝕜 s x) fderivWithin_of_mem_nhdsWithin (nhdsWithin_mono _ st self_mem_nhdsWithin) ht h theorem fderivWithin_inter (ht : t ∈ 𝓝 x) : fderivWithin 𝕜 f (s ∩ t) x = fderivWithin 𝕜 f s x := by + classical simp [fderivWithin, hasFDerivWithinAt_inter ht, DifferentiableWithinAt] theorem fderivWithin_of_mem_nhds (h : s ∈ 𝓝 x) : fderivWithin 𝕜 f s x = fderiv 𝕜 f x := by @@ -767,6 +766,7 @@ theorem differentiableWithinAt_congr_set (h : s =ᶠ[𝓝 x] t) : theorem fderivWithin_congr_set' (y : E) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : fderivWithin 𝕜 f s x = fderivWithin 𝕜 f t x := by + classical simp only [fderivWithin, differentiableWithinAt_congr_set' _ h, hasFDerivWithinAt_congr_set' _ h] theorem fderivWithin_congr_set (h : s =ᶠ[𝓝 x] t) : fderivWithin 𝕜 f s x = fderivWithin 𝕜 f t x := @@ -897,6 +897,7 @@ theorem DifferentiableWithinAt.fderivWithin_congr_mono (h : DifferentiableWithin theorem Filter.EventuallyEq.fderivWithin_eq (hs : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : fderivWithin 𝕜 f₁ s x = fderivWithin 𝕜 f s x := by + classical simp only [fderivWithin, DifferentiableWithinAt, hs.hasFDerivWithinAt_iff hx] theorem Filter.EventuallyEq.fderivWithin_eq_of_mem (hs : f₁ =ᶠ[𝓝[s] x] f) (hx : x ∈ s) : diff --git a/Mathlib/Analysis/Calculus/FDeriv/Comp.lean b/Mathlib/Analysis/Calculus/FDeriv/Comp.lean index df6c6a6a2be03..64aaa3ac3b12b 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Comp.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Comp.lean @@ -16,10 +16,7 @@ composition of functions (the chain rule). -/ -open Filter Asymptotics ContinuousLinearMap Set Metric - -open scoped Classical -open Topology NNReal Filter Asymptotics ENNReal +open Filter Asymptotics ContinuousLinearMap Set Metric Topology NNReal ENNReal noncomputable section diff --git a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean index f956fb2da11a4..b859dc97ba096 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean @@ -20,10 +20,7 @@ We also prove the usual formula for the derivative of the inverse function, assu The inverse function theorem is in `Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean`. -/ -open Filter Asymptotics ContinuousLinearMap Set Metric - -open scoped Classical -open Topology NNReal Filter Asymptotics ENNReal +open Filter Asymptotics ContinuousLinearMap Set Metric Topology NNReal ENNReal noncomputable section diff --git a/Mathlib/Analysis/Calculus/FDeriv/Prod.lean b/Mathlib/Analysis/Calculus/FDeriv/Prod.lean index f1657e4a5b59f..c5e33c862226a 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Prod.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Prod.lean @@ -17,10 +17,7 @@ cartesian products of functions, and functions into Pi-types. -/ -open Filter Asymptotics ContinuousLinearMap Set Metric - -open scoped Classical -open Topology NNReal Filter Asymptotics ENNReal +open Filter Asymptotics ContinuousLinearMap Set Metric Topology NNReal ENNReal noncomputable section diff --git a/Mathlib/Analysis/Calculus/FDeriv/RestrictScalars.lean b/Mathlib/Analysis/Calculus/FDeriv/RestrictScalars.lean index 7936b7024558c..cac15836bc67a 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/RestrictScalars.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/RestrictScalars.lean @@ -16,10 +16,7 @@ the scalar restriction of a linear map. -/ -open Filter Asymptotics ContinuousLinearMap Set Metric - -open scoped Classical -open Topology NNReal Filter Asymptotics ENNReal +open Filter Asymptotics ContinuousLinearMap Set Metric Topology NNReal ENNReal noncomputable section diff --git a/Mathlib/Analysis/Calculus/FDeriv/Star.lean b/Mathlib/Analysis/Calculus/FDeriv/Star.lean index 933c9c2c4fdc1..363f57e88b8a7 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Star.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Star.lean @@ -20,8 +20,6 @@ star operation; which as should be expected rules out `𝕜 = ℂ`. -/ -open scoped Classical - variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] [StarRing 𝕜] [TrivialStar 𝕜] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] variable {F : Type*} [NormedAddCommGroup F] [StarAddMonoid F] [NormedSpace 𝕜 F] [StarModule 𝕜 F] diff --git a/Mathlib/Analysis/Calculus/LogDeriv.lean b/Mathlib/Analysis/Calculus/LogDeriv.lean index 77181f252bb0a..f62191529939f 100644 --- a/Mathlib/Analysis/Calculus/LogDeriv.lean +++ b/Mathlib/Analysis/Calculus/LogDeriv.lean @@ -17,7 +17,7 @@ noncomputable section open Filter Function -open scoped Topology Classical +open scoped Topology variable {𝕜 𝕜': Type*} [NontriviallyNormedField 𝕜] [NontriviallyNormedField 𝕜'] [NormedAlgebra 𝕜 𝕜'] diff --git a/Mathlib/Analysis/Convex/Combination.lean b/Mathlib/Analysis/Convex/Combination.lean index 983f2d7608437..5e09377713d95 100644 --- a/Mathlib/Analysis/Convex/Combination.lean +++ b/Mathlib/Analysis/Convex/Combination.lean @@ -24,10 +24,7 @@ lemmas unconditional on the sum of the weights being `1`. -/ -open Set Function - -open scoped Classical -open Pointwise +open Set Function Pointwise universe u u' @@ -48,6 +45,7 @@ open Finset theorem Finset.centerMass_empty : (∅ : Finset ι).centerMass w z = 0 := by simp only [centerMass, sum_empty, smul_zero] +open scoped Classical in theorem Finset.centerMass_pair (hne : i ≠ j) : ({i, j} : Finset ι).centerMass w z = (w i / (w i + w j)) • z i + (w j / (w i + w j)) • z j := by simp only [centerMass, sum_pair hne] @@ -55,6 +53,7 @@ theorem Finset.centerMass_pair (hne : i ≠ j) : variable {w} +open scoped Classical in theorem Finset.centerMass_insert (ha : i ∉ t) (hw : ∑ j ∈ t, w j ≠ 0) : (insert i t).centerMass w z = (w i / (w i + ∑ j ∈ t, w j)) • z i + @@ -105,6 +104,7 @@ theorem Finset.centerMass_segment (s : Finset ι) (w₁ w₂ : ι → R) (z : ι simp only [Finset.centerMass_eq_of_sum_1, Finset.centerMass_eq_of_sum_1 _ _ hw, smul_sum, sum_add_distrib, add_smul, mul_smul, *] +open scoped Classical in theorem Finset.centerMass_ite_eq (hi : i ∈ t) : t.centerMass (fun j => if i = j then (1 : R) else 0) z = z i := by rw [Finset.centerMass_eq_of_sum_1] @@ -154,6 +154,7 @@ lemma Finset.centerMass_of_sum_add_sum_eq_zero {s t : Finset ι} provided that all weights are non-negative, and the total weight is positive. -/ theorem Convex.centerMass_mem (hs : Convex R s) : (∀ i ∈ t, 0 ≤ w i) → (0 < ∑ i ∈ t, w i) → (∀ i ∈ t, z i ∈ s) → t.centerMass w z ∈ s := by + classical induction' t using Finset.induction with i t hi ht · simp [lt_irrefl] intro h₀ hpos hmem @@ -199,6 +200,7 @@ theorem Convex.finsum_mem {ι : Sort*} {w : ι → R} {z : ι → E} {s : Set E} theorem convex_iff_sum_mem : Convex R s ↔ ∀ (t : Finset E) (w : E → R), (∀ i ∈ t, 0 ≤ w i) → ∑ i ∈ t, w i = 1 → (∀ x ∈ t, x ∈ s) → (∑ x ∈ t, w x • x) ∈ s := by + classical refine ⟨fun hs t w hw₀ hw₁ hts => hs.sum_mem hw₀ hw₁ hts, ?_⟩ intro h x hx y hy a b ha hb hab by_cases h_cases : x = y @@ -270,6 +272,7 @@ theorem Finset.centroid_mem_convexHull (s : Finset E) (hs : s.Nonempty) : theorem convexHull_range_eq_exists_affineCombination (v : ι → E) : convexHull R (range v) = { x | ∃ (s : Finset ι) (w : ι → R), (∀ i ∈ s, 0 ≤ w i) ∧ s.sum w = 1 ∧ s.affineCombination R v w = x } := by + classical refine Subset.antisymm (convexHull_min ?_ ?_) ?_ · intro x hx obtain ⟨i, hi⟩ := Set.mem_range.mp hx @@ -351,6 +354,7 @@ lemma mem_convexHull_iff_exists_fintype {s : Set E} {x : E} : theorem Finset.convexHull_eq (s : Finset E) : convexHull R ↑s = { x : E | ∃ w : E → R, (∀ y ∈ s, 0 ≤ w y) ∧ ∑ y ∈ s, w y = 1 ∧ s.centerMass w id = x } := by + classical refine Set.Subset.antisymm (convexHull_min ?_ ?_) ?_ · intro x hx rw [Finset.mem_coe] at hx @@ -390,6 +394,7 @@ theorem Set.Finite.convexHull_eq {s : Set E} (hs : s.Finite) : convexHull R s = /-- A weak version of Carathéodory's theorem. -/ theorem convexHull_eq_union_convexHull_finite_subsets (s : Set E) : convexHull R s = ⋃ (t : Finset E) (_ : ↑t ⊆ s), convexHull R ↑t := by + classical refine Subset.antisymm ?_ ?_ · rw [_root_.convexHull_eq] rintro x ⟨ι, t, w, z, hw₀, hw₁, hz, rfl⟩ @@ -458,6 +463,7 @@ theorem convexHull_sum {ι} (s : Finset ι) (t : ι → Set E) : variable (ι) [Fintype ι] {f : ι → R} +open scoped Classical in /-- `stdSimplex 𝕜 ι` is the convex hull of the canonical basis in `ι → 𝕜`. -/ theorem convexHull_basis_eq_stdSimplex : convexHull R (range fun i j : ι => if i = j then (1 : R) else 0) = stdSimplex R ι := by @@ -480,6 +486,7 @@ to prove that this map is linear. -/ theorem Set.Finite.convexHull_eq_image {s : Set E} (hs : s.Finite) : convexHull R s = haveI := hs.fintype (⇑(∑ x : s, (@LinearMap.proj R s _ (fun _ => R) _ _ x).smulRight x.1)) '' stdSimplex R s := by + classical letI := hs.fintype rw [← convexHull_basis_eq_stdSimplex, LinearMap.image_convexHull, ← Set.range_comp] apply congr_arg @@ -519,6 +526,7 @@ 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 + classical 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'] @@ -540,6 +548,7 @@ lemma AffineIndependent.convexHull_inter (hs : AffineIndependent R ((↑) : s simp_intro hx₁ hx₂ simp [ht x hx₁ hx₂] +open scoped Classical in /-- Two simplices glue nicely if the union of their vertices is affine independent. Note that `AffineIndependent.convexHull_inter` should be more versatile in most use cases. -/ @@ -556,6 +565,7 @@ variable {𝕜 ι : Type*} {E : ι → Type*} [Finite ι] [LinearOrderedField open Finset Fintype lemma mem_convexHull_pi (h : ∀ i ∈ s, x i ∈ convexHull 𝕜 (t i)) : x ∈ convexHull 𝕜 (s.pi t) := by + classical cases nonempty_fintype ι wlog hs : s = Set.univ generalizing s t · rw [← pi_univ_ite] diff --git a/Mathlib/Analysis/Convex/Cone/Basic.lean b/Mathlib/Analysis/Convex/Cone/Basic.lean index 303a1a919f4f2..e85e4d156f432 100644 --- a/Mathlib/Analysis/Convex/Cone/Basic.lean +++ b/Mathlib/Analysis/Convex/Cone/Basic.lean @@ -41,10 +41,7 @@ assert_not_exists NormedSpace assert_not_exists Real assert_not_exists Cardinal -open Set LinearMap - -open scoped Classical -open Pointwise +open Set LinearMap Pointwise variable {𝕜 E F G : Type*} diff --git a/Mathlib/Analysis/Convex/Cone/InnerDual.lean b/Mathlib/Analysis/Convex/Cone/InnerDual.lean index 466d181b61ba0..e77cf62b4dc52 100644 --- a/Mathlib/Analysis/Convex/Cone/InnerDual.lean +++ b/Mathlib/Analysis/Convex/Cone/InnerDual.lean @@ -27,10 +27,7 @@ We prove the following theorems: [Farkas lemma](https://en.wikipedia.org/wiki/Farkas%27_lemma#Geometric_interpretation). -/ -open Set LinearMap - -open scoped Classical -open Pointwise +open Set LinearMap Pointwise /-! ### The dual cone -/ diff --git a/Mathlib/Analysis/Convex/Exposed.lean b/Mathlib/Analysis/Convex/Exposed.lean index c3d804c5b270b..c419ff008281f 100644 --- a/Mathlib/Analysis/Convex/Exposed.lean +++ b/Mathlib/Analysis/Convex/Exposed.lean @@ -39,11 +39,7 @@ See chapter 8 of [Barry Simon, *Convexity*][simon2011] Prove lemmas relating exposed sets and points to the intrinsic frontier. -/ - -open scoped Classical -open Affine - -open Set +open Affine Set section PreorderSemiring @@ -140,6 +136,7 @@ protected theorem inter [ContinuousAdd 𝕜] {A B C : Set E} (hB : IsExposed theorem sInter [ContinuousAdd 𝕜] {F : Finset (Set E)} (hF : F.Nonempty) (hAF : ∀ B ∈ F, IsExposed 𝕜 A B) : IsExposed 𝕜 A (⋂₀ F) := by + classical induction F using Finset.induction with | empty => exfalso; exact Finset.not_nonempty_empty hF | @insert C F _ hF' => diff --git a/Mathlib/Analysis/Convex/Extrema.lean b/Mathlib/Analysis/Convex/Extrema.lean index 798b67c20f4f8..91f73a294be07 100644 --- a/Mathlib/Analysis/Convex/Extrema.lean +++ b/Mathlib/Analysis/Convex/Extrema.lean @@ -19,10 +19,7 @@ a global minimum, and likewise for concave functions. variable {E β : Type*} [AddCommGroup E] [TopologicalSpace E] [Module ℝ E] [TopologicalAddGroup E] [ContinuousSMul ℝ E] [OrderedAddCommGroup β] [Module ℝ β] [OrderedSMul ℝ β] {s : Set E} -open Set Filter Function - -open scoped Classical -open Topology +open Set Filter Function Topology /-- Helper lemma for the more general case: `IsMinOn.of_isLocalMinOn_of_convexOn`. -/ diff --git a/Mathlib/Analysis/Convex/Extreme.lean b/Mathlib/Analysis/Convex/Extreme.lean index 652e14f30bcc9..c08e5824b54eb 100644 --- a/Mathlib/Analysis/Convex/Extreme.lean +++ b/Mathlib/Analysis/Convex/Extreme.lean @@ -39,10 +39,7 @@ Prove lemmas relating extreme sets and points to the intrinsic frontier. -/ -open Function Set - -open scoped Classical -open Affine +open Function Set Affine variable {𝕜 E F ι : Type*} {π : ι → Type*} @@ -187,6 +184,7 @@ theorem extremePoints_prod (s : Set E) (t : Set F) : @[simp] theorem extremePoints_pi (s : ∀ i, Set (π i)) : (univ.pi s).extremePoints 𝕜 = univ.pi fun i ↦ (s i).extremePoints 𝕜 := by + classical ext x simp only [mem_extremePoints, mem_pi, mem_univ, true_imp_iff, @forall_and ι] refine and_congr_right fun hx ↦ ⟨fun h i ↦ ?_, fun h ↦ ?_⟩ diff --git a/Mathlib/Analysis/Convex/Function.lean b/Mathlib/Analysis/Convex/Function.lean index d3bd49537ee13..8ff7dc2396895 100644 --- a/Mathlib/Analysis/Convex/Function.lean +++ b/Mathlib/Analysis/Convex/Function.lean @@ -26,8 +26,6 @@ a convex set. * `StrictConcaveOn 𝕜 s f`: The function `f` is strictly concave on `s` with scalars `𝕜`. -/ - -open scoped Classical open LinearMap Set Convex Pointwise variable {𝕜 E F α β ι : Type*} diff --git a/Mathlib/Analysis/Convex/Independent.lean b/Mathlib/Analysis/Convex/Independent.lean index 9547830a35683..694644a35ed21 100644 --- a/Mathlib/Analysis/Convex/Independent.lean +++ b/Mathlib/Analysis/Convex/Independent.lean @@ -40,10 +40,7 @@ independence, convex position -/ -open scoped Classical -open Affine - -open Finset Function +open Affine Finset Function variable {𝕜 E ι : Type*} @@ -158,6 +155,7 @@ section LinearOrderedField variable [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} +open scoped Classical in /-- To check convex independence, one only has to check finsets thanks to Carathéodory's theorem. -/ theorem convexIndependent_iff_finset {p : ι → E} : ConvexIndependent 𝕜 p ↔ diff --git a/Mathlib/Analysis/Convex/Jensen.lean b/Mathlib/Analysis/Convex/Jensen.lean index 0fb76f66c1e58..618b754855b53 100644 --- a/Mathlib/Analysis/Convex/Jensen.lean +++ b/Mathlib/Analysis/Convex/Jensen.lean @@ -31,10 +31,7 @@ As corollaries, we get: -/ -open Finset LinearMap Set - -open scoped Classical -open Convex Pointwise +open Finset LinearMap Set Convex Pointwise variable {𝕜 E F β ι : Type*} diff --git a/Mathlib/Analysis/Convex/KreinMilman.lean b/Mathlib/Analysis/Convex/KreinMilman.lean index d5a0fa7d21ade..9b3f6f3ff9d87 100644 --- a/Mathlib/Analysis/Convex/KreinMilman.lean +++ b/Mathlib/Analysis/Convex/KreinMilman.lean @@ -51,7 +51,6 @@ See chapter 8 of [Barry Simon, *Convexity*][simon2011] -/ open Set -open scoped Classical variable {E F : Type*} [AddCommGroup E] [Module ℝ E] [TopologicalSpace E] [T2Space E] [TopologicalAddGroup E] [ContinuousSMul ℝ E] [LocallyConvexSpace ℝ E] {s : Set E} diff --git a/Mathlib/Analysis/InnerProductSpace/Dual.lean b/Mathlib/Analysis/InnerProductSpace/Dual.lean index 5fa801d055b82..80077cbfd8a9d 100644 --- a/Mathlib/Analysis/InnerProductSpace/Dual.lean +++ b/Mathlib/Analysis/InnerProductSpace/Dual.lean @@ -36,7 +36,6 @@ dual, Fréchet-Riesz noncomputable section -open scoped Classical open ComplexConjugate universe u v diff --git a/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean b/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean index 0becfb9ef8a29..e80e7b0443eb7 100644 --- a/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean +++ b/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean @@ -50,7 +50,7 @@ noncomputable section open RCLike -open scoped ComplexConjugate Classical +open scoped ComplexConjugate variable {𝕜 E F : Type*} [RCLike 𝕜] variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] @@ -142,6 +142,7 @@ theorem adjointAux_unique (y : T.adjointDomain) {x₀ : E} variable (T) +open scoped Classical in /-- The adjoint operator as a partially defined linear operator. -/ def adjoint : F →ₗ.[𝕜] E where domain := T.adjointDomain @@ -163,10 +164,12 @@ theorem mem_adjoint_domain_of_exists (y : F) (h : ∃ w : E, ∀ x : T.domain, exact funext fun x => (hw x).symm theorem adjoint_apply_of_not_dense (hT : ¬Dense (T.domain : Set E)) (y : T†.domain) : T† y = 0 := by + classical change (if hT : Dense (T.domain : Set E) then adjointAux hT else 0) y = _ simp only [hT, not_false_iff, dif_neg, LinearMap.zero_apply] theorem adjoint_apply_of_dense (y : T†.domain) : T† y = adjointAux hT y := by + classical change (if hT : Dense (T.domain : Set E) then adjointAux hT else 0) y = _ simp only [hT, dif_pos, LinearMap.coe_mk] diff --git a/Mathlib/Analysis/Normed/Module/Dual.lean b/Mathlib/Analysis/Normed/Module/Dual.lean index 99cc8edec9dd3..8a947e649a0b2 100644 --- a/Mathlib/Analysis/Normed/Module/Dual.lean +++ b/Mathlib/Analysis/Normed/Module/Dual.lean @@ -41,7 +41,6 @@ dual, polar noncomputable section -open scoped Classical open Topology Bornology universe u v diff --git a/Mathlib/Analysis/Normed/Operator/Banach.lean b/Mathlib/Analysis/Normed/Operator/Banach.lean index 27ca8d21c6ddf..6df43033c34c9 100644 --- a/Mathlib/Analysis/Normed/Operator/Banach.lean +++ b/Mathlib/Analysis/Normed/Operator/Banach.lean @@ -16,7 +16,6 @@ This file contains the Banach open mapping theorem, i.e., the fact that a biject bounded linear map between Banach spaces has a bounded inverse. -/ -open scoped Classical open Function Metric Set Filter Finset Topology NNReal open LinearMap (range ker) diff --git a/Mathlib/Analysis/Normed/Ring/Units.lean b/Mathlib/Analysis/Normed/Ring/Units.lean index e0b887991a865..88bb4dad29170 100644 --- a/Mathlib/Analysis/Normed/Ring/Units.lean +++ b/Mathlib/Analysis/Normed/Ring/Units.lean @@ -88,8 +88,6 @@ end nonunits namespace NormedRing -open scoped Classical - open Asymptotics Filter Metric Finset Ring theorem inverse_one_sub (t : R) (h : ‖t‖ < 1) : inverse (1 - t) = ↑(Units.oneSub t h)⁻¹ := by diff --git a/Mathlib/Analysis/RCLike/Lemmas.lean b/Mathlib/Analysis/RCLike/Lemmas.lean index 9ce1f0f40f0a5..2fa56df7b41a0 100644 --- a/Mathlib/Analysis/RCLike/Lemmas.lean +++ b/Mathlib/Analysis/RCLike/Lemmas.lean @@ -20,8 +20,6 @@ end Polynomial namespace FiniteDimensional -open scoped Classical - open RCLike library_note "RCLike instance"/-- diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean index 3dc511fcfb8d0..e6339d7d7c234 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean @@ -17,7 +17,6 @@ located here. noncomputable section -open scoped Classical open Real Topology NNReal ENNReal Filter ComplexConjugate Finset Set /-! diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean index c9a10865fd507..49e40ea604b12 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean @@ -11,7 +11,6 @@ import Mathlib.Analysis.SpecialFunctions.Complex.Log We construct the power functions `x ^ y`, where `x` and `y` are complex numbers. -/ -open scoped Classical open Real Topology Filter ComplexConjugate Finset Set namespace Complex diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean index 63668985e94b3..21374aff28eb8 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean @@ -15,10 +15,7 @@ This file contains lemmas about continuity of the power functions on `ℂ`, `ℝ noncomputable section -open scoped Classical -open Real Topology NNReal ENNReal Filter ComplexConjugate - -open Filter Finset Set +open Real Topology NNReal ENNReal Filter ComplexConjugate Finset Set section CpowLimits diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean index acdcc41b5be97..43bf1bcc6be45 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean @@ -18,10 +18,7 @@ We also prove basic properties of these functions. noncomputable section -open scoped Classical -open Real NNReal ENNReal ComplexConjugate - -open Finset Function Set +open Real NNReal ENNReal ComplexConjugate Finset Function Set namespace NNReal variable {x : ℝ≥0} {w y z : ℝ} @@ -680,6 +677,7 @@ theorem coe_mul_rpow (x y : ℝ≥0) (z : ℝ) : ((x : ℝ≥0∞) * y) ^ z = (x theorem prod_coe_rpow {ι} (s : Finset ι) (f : ι → ℝ≥0) (r : ℝ) : ∏ i ∈ s, (f i : ℝ≥0∞) ^ r = ((∏ i ∈ s, f i : ℝ≥0) : ℝ≥0∞) ^ r := by + classical induction s using Finset.induction with | empty => simp | insert hi ih => simp_rw [prod_insert hi, ih, ← coe_mul_rpow, coe_mul] @@ -692,6 +690,7 @@ theorem mul_rpow_of_nonneg (x y : ℝ≥0∞) {z : ℝ} (hz : 0 ≤ z) : (x * y) theorem prod_rpow_of_ne_top {ι} {s : Finset ι} {f : ι → ℝ≥0∞} (hf : ∀ i ∈ s, f i ≠ ∞) (r : ℝ) : ∏ i ∈ s, f i ^ r = (∏ i ∈ s, f i) ^ r := by + classical induction s using Finset.induction with | empty => simp | @insert i s hi ih => @@ -701,6 +700,7 @@ theorem prod_rpow_of_ne_top {ι} {s : Finset ι} {f : ι → ℝ≥0∞} (hf : theorem prod_rpow_of_nonneg {ι} {s : Finset ι} {f : ι → ℝ≥0∞} {r : ℝ} (hr : 0 ≤ r) : ∏ i ∈ s, f i ^ r = (∏ i ∈ s, f i) ^ r := by + classical induction s using Finset.induction with | empty => simp | insert hi ih => simp_rw [prod_insert hi, ih, ← mul_rpow_of_nonneg _ _ hr] diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index 0a7f64d0a15aa..809b94d5aa7d0 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -15,10 +15,7 @@ We construct the power functions `x ^ y`, where `x` and `y` are real numbers. noncomputable section -open scoped Classical -open Real ComplexConjugate - -open Finset Set +open Real ComplexConjugate Finset Set /- ## Definitions diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index ff1dbf10a31ac..9caf00dfe93d5 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -44,7 +44,6 @@ sin, cos, tan, angle noncomputable section -open scoped Classical open Topology Filter Set namespace Complex diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Inverse.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Inverse.lean index 07e897502b584..86494264b3012 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Inverse.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Inverse.lean @@ -18,12 +18,7 @@ Basic inequalities on trigonometric functions. noncomputable section -open scoped Classical -open Topology Filter - -open Set Filter - -open Real +open Topology Filter Set Filter Real namespace Real variable {x y : ℝ} diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index d86c021d2cfec..c226f8f39d7af 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -21,11 +21,7 @@ instances of these such as `ℝ`, `ℝ≥0` and `ℝ≥0∞`. noncomputable section -open scoped Classical -open Set Function Filter Finset Metric - -open scoped Classical -open Topology Nat uniformity NNReal ENNReal +open Set Function Filter Finset Metric Topology Nat uniformity NNReal ENNReal variable {α : Type*} {β : Type*} {ι : Type*} @@ -552,6 +548,7 @@ def posSumOfEncodable {ε : ℝ} (hε : 0 < ε) (ι) [Encodable ι] : theorem Set.Countable.exists_pos_hasSum_le {ι : Type*} {s : Set ι} (hs : s.Countable) {ε : ℝ} (hε : 0 < ε) : ∃ ε' : ι → ℝ, (∀ i, 0 < ε' i) ∧ ∃ c, HasSum (fun i : s ↦ ε' i) c ∧ c ≤ ε := by + classical haveI := hs.toEncodable rcases posSumOfEncodable hε s with ⟨f, hf0, ⟨c, hfc, hcε⟩⟩ refine ⟨fun i ↦ if h : i ∈ s then f ⟨i, h⟩ else 1, fun i ↦ ?_, ⟨c, ?_, hcε⟩⟩ @@ -563,6 +560,7 @@ theorem Set.Countable.exists_pos_hasSum_le {ι : Type*} {s : Set ι} (hs : s.Cou theorem Set.Countable.exists_pos_forall_sum_le {ι : Type*} {s : Set ι} (hs : s.Countable) {ε : ℝ} (hε : 0 < ε) : ∃ ε' : ι → ℝ, (∀ i, 0 < ε' i) ∧ ∀ t : Finset ι, ↑t ⊆ s → ∑ i ∈ t, ε' i ≤ ε := by + classical rcases hs.exists_pos_hasSum_le hε with ⟨ε', hpos, c, hε'c, hcε⟩ refine ⟨ε', hpos, fun t ht ↦ ?_⟩ rw [← sum_subtype_of_mem _ ht] diff --git a/Mathlib/CategoryTheory/Countable.lean b/Mathlib/CategoryTheory/Countable.lean index e66571ff75dae..55743e7cf997d 100644 --- a/Mathlib/CategoryTheory/Countable.lean +++ b/Mathlib/CategoryTheory/Countable.lean @@ -16,8 +16,6 @@ A category is countable in this sense if it has countably many objects and count universe w v u -open scoped Classical - noncomputable section namespace CategoryTheory diff --git a/Mathlib/CategoryTheory/FinCategory/AsType.lean b/Mathlib/CategoryTheory/FinCategory/AsType.lean index dafd99f417b23..ad3fd1ca2e32b 100644 --- a/Mathlib/CategoryTheory/FinCategory/AsType.lean +++ b/Mathlib/CategoryTheory/FinCategory/AsType.lean @@ -13,8 +13,6 @@ import Mathlib.CategoryTheory.FinCategory.Basic universe w v u -open scoped Classical - noncomputable section namespace CategoryTheory diff --git a/Mathlib/CategoryTheory/FinCategory/Basic.lean b/Mathlib/CategoryTheory/FinCategory/Basic.lean index 6e3d5cbfa8f82..548e25f7b10c3 100644 --- a/Mathlib/CategoryTheory/FinCategory/Basic.lean +++ b/Mathlib/CategoryTheory/FinCategory/Basic.lean @@ -23,8 +23,6 @@ having to supply instances or delay with non-defeq conflicts between instances. universe w v u -open scoped Classical - noncomputable section namespace CategoryTheory @@ -33,6 +31,7 @@ instance discreteFintype {α : Type*} [Fintype α] : Fintype (Discrete α) := Fintype.ofEquiv α discreteEquiv.symm instance discreteHomFintype {α : Type*} (X Y : Discrete α) : Fintype (X ⟶ Y) := by + classical apply ULift.fintype /-- A category with a `Fintype` of objects, and a `Fintype` for each morphism space. -/ diff --git a/Mathlib/CategoryTheory/FintypeCat.lean b/Mathlib/CategoryTheory/FintypeCat.lean index abe6c3d23a130..871779a37dcf5 100644 --- a/Mathlib/CategoryTheory/FintypeCat.lean +++ b/Mathlib/CategoryTheory/FintypeCat.lean @@ -21,9 +21,6 @@ are `Fin n` for `n : ℕ`. We prove that the obvious inclusion functor We prove that `FintypeCat.Skeleton` is a skeleton of `FintypeCat` in `FintypeCat.isSkeleton`. -/ - -open scoped Classical - open CategoryTheory /-- The category of finite types. -/ diff --git a/Mathlib/CategoryTheory/GlueData.lean b/Mathlib/CategoryTheory/GlueData.lean index 5f09d792ec41d..2c0b603d6f43f 100644 --- a/Mathlib/CategoryTheory/GlueData.lean +++ b/Mathlib/CategoryTheory/GlueData.lean @@ -394,8 +394,7 @@ attribute [reassoc (attr := simp)] GlueData'.t_inv GlueData'.cocycle variable {C} -open scoped Classical - +open scoped Classical in /-- (Implementation detail) the constructed `GlueData.f` from a `GlueData'`. -/ abbrev GlueData'.f' (D : GlueData' C) (i j : D.J) : (if h : i = j then D.U i else D.V i j h) ⟶ D.U i := @@ -421,6 +420,7 @@ instance (D : GlueData' C) (i j k : D.J) : have {X Y Z : C} (f : X ⟶ Y) (e : Z = X) : HEq (eqToHom e ≫ f) f := by subst e; simp convert D.f_hasPullback i j k hij hik <;> simp [GlueData'.f', hij, hik, this] +open scoped Classical in /-- (Implementation detail) the constructed `GlueData.t'` from a `GlueData'`. -/ def GlueData'.t'' (D : GlueData' C) (i j k : D.J) : pullback (D.f' i j) (D.f' i k) ⟶ pullback (D.f' j k) (D.f' j i) := @@ -448,6 +448,7 @@ def GlueData'.t'' (D : GlueData' C) (i j k : D.J) : pullback.map _ _ _ _ (eqToHom (by aesop)) (eqToHom (by aesop)) (eqToHom (by aesop)) (by delta f'; aesop) (by delta f'; aesop) +open scoped Classical in /-- The constructed `GlueData` of a `GlueData'`, where `GlueData'` is a variant of `GlueData` that only requires conditions on `V (i, j)` when `i ≠ j`. diff --git a/Mathlib/CategoryTheory/Limits/Bicones.lean b/Mathlib/CategoryTheory/Limits/Bicones.lean index a2be62de66858..9def7f73005be 100644 --- a/Mathlib/CategoryTheory/Limits/Bicones.lean +++ b/Mathlib/CategoryTheory/Limits/Bicones.lean @@ -27,8 +27,6 @@ noncomputable section open CategoryTheory.Limits -open scoped Classical - namespace CategoryTheory section Bicone @@ -45,6 +43,7 @@ variable (J : Type u₁) instance : Inhabited (Bicone J) := ⟨Bicone.left⟩ +open scoped Classical in instance finBicone [Fintype J] : Fintype (Bicone J) where elems := [Bicone.left, Bicone.right].toFinset ∪ Finset.image Bicone.diagram Fintype.elems complete j := by @@ -64,7 +63,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 only [diagram.injEq] <;> infer_instance + classical cases f <;> cases g <;> simp only [diagram.injEq] <;> infer_instance @[simps] instance biconeCategoryStruct : CategoryStruct (Bicone J) where @@ -116,6 +115,7 @@ def biconeMk {C : Type u₁} [Category.{v₁} C] {F : J ⥤ C} (c₁ c₂ : Cone · cases g apply F.map_comp +open scoped Classical in instance finBiconeHom [FinCategory J] (j k : Bicone J) : Fintype (j ⟶ k) := by cases j <;> cases k · exact diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index 29f15f3221a2d..cf91656226688 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -52,8 +52,6 @@ open CategoryTheory open CategoryTheory.Functor -open scoped Classical - namespace CategoryTheory namespace Limits @@ -63,6 +61,7 @@ universe uC' uC uD' uD variable {C : Type uC} [Category.{uC'} C] [HasZeroMorphisms C] variable {D : Type uD} [Category.{uD'} D] [HasZeroMorphisms D] +open scoped Classical in /-- A `c : Bicone F` is: * an object `c.pt` and * morphisms `π j : pt ⟶ F j` and `ι j : F j ⟶ pt` for each `j`, @@ -212,6 +211,7 @@ theorem toCocone_inj (B : Bicone F) (j : J) : Cofan.inj B.toCocone j = B.ι j := theorem toCocone_ι_app_mk (B : Bicone F) (j : J) : B.toCocone.ι.app ⟨j⟩ = B.ι j := rfl +open scoped Classical in /-- We can turn any limit cone over a discrete collection of objects into a bicone. -/ @[simps] def ofLimitCone {f : J → C} {t : Cone (Discrete.functor f)} (ht : IsLimit t) : Bicone f where @@ -220,12 +220,14 @@ def ofLimitCone {f : J → C} {t : Cone (Discrete.functor f)} (ht : IsLimit t) : ι j := ht.lift (Fan.mk _ fun j' => if h : j = j' then eqToHom (congr_arg f h) else 0) ι_π j j' := by simp +open scoped Classical in theorem ι_of_isLimit {f : J → C} {t : Bicone f} (ht : IsLimit t.toCone) (j : J) : t.ι j = ht.lift (Fan.mk _ fun j' => if h : j = j' then eqToHom (congr_arg f h) else 0) := ht.hom_ext fun j' => by rw [ht.fac] simp [t.ι_π] +open scoped Classical in /-- We can turn any colimit cocone over a discrete collection of objects into a bicone. -/ @[simps] def ofColimitCocone {f : J → C} {t : Cocone (Discrete.functor f)} (ht : IsColimit t) : @@ -235,6 +237,7 @@ def ofColimitCocone {f : J → C} {t : Cocone (Discrete.functor f)} (ht : IsColi ι j := t.ι.app ⟨j⟩ ι_π j j' := by simp +open scoped Classical in theorem π_of_isColimit {f : J → C} {t : Bicone f} (ht : IsColimit t.toCocone) (j : J) : t.π j = ht.desc (Cofan.mk _ fun j' => if h : j' = j then eqToHom (congr_arg f h) else 0) := ht.hom_ext fun j' => by @@ -561,11 +564,13 @@ def HasBiproductsOfShape.colimIsoLim [HasBiproductsOfShape J C] : NatIso.ofComponents (fun F => (Sigma.isoColimit F).symm ≪≫ (biproduct.isoCoproduct _).symm ≪≫ biproduct.isoProduct _ ≪≫ Pi.isoLimit F) fun η => colimit.hom_ext fun ⟨i⟩ => limit.hom_ext fun ⟨j⟩ => by + classical by_cases h : i = j <;> simp_all [h, Sigma.isoColimit, Pi.isoLimit, biproduct.ι_π, biproduct.ι_π_assoc] theorem biproduct.map_eq_map' {f g : J → C} [HasBiproduct f] [HasBiproduct g] (p : ∀ b, f b ⟶ g b) : biproduct.map p = biproduct.map' p := by + classical ext dsimp simp only [Discrete.natTrans_app, Limits.IsColimit.ι_map_assoc, Limits.IsLimit.map_π, @@ -612,6 +617,7 @@ def biproduct.mapIso {f g : J → C} [HasBiproduct f] [HasBiproduct g] (p : ∀ instance biproduct.map_epi {f g : J → C} [HasBiproduct f] [HasBiproduct g] (p : ∀ j, f j ⟶ g j) [∀ j, Epi (p j)] : Epi (biproduct.map p) := by + classical have : biproduct.map p = (biproduct.isoCoproduct _).hom ≫ Sigma.map p ≫ (biproduct.isoCoproduct _).inv := by ext @@ -742,6 +748,7 @@ def biproduct.toSubtype : ⨁ f ⟶ ⨁ Subtype.restrict p f := theorem biproduct.fromSubtype_π [DecidablePred p] (j : J) : biproduct.fromSubtype f p ≫ biproduct.π f j = if h : p j then biproduct.π (Subtype.restrict p f) ⟨j, h⟩ else 0 := by + classical ext i; dsimp rw [biproduct.fromSubtype, biproduct.ι_desc_assoc, biproduct.ι_π] by_cases h : p j @@ -758,6 +765,7 @@ theorem biproduct.fromSubtype_eq_lift [DecidablePred p] : @[reassoc] -- Porting note: both version solved using simp theorem biproduct.fromSubtype_π_subtype (j : Subtype p) : biproduct.fromSubtype f p ≫ biproduct.π f j = biproduct.π (Subtype.restrict p f) j := by + classical ext rw [biproduct.fromSubtype, biproduct.ι_desc_assoc, biproduct.ι_π, biproduct.ι_π] split_ifs with h₁ h₂ h₂ @@ -772,6 +780,7 @@ theorem biproduct.toSubtype_π (j : Subtype p) : theorem biproduct.ι_toSubtype [DecidablePred p] (j : J) : biproduct.ι f j ≫ biproduct.toSubtype f p = if h : p j then biproduct.ι (Subtype.restrict p f) ⟨j, h⟩ else 0 := by + classical ext i rw [biproduct.toSubtype, Category.assoc, biproduct.lift_π, biproduct.ι_π] by_cases h : p j @@ -788,6 +797,7 @@ theorem biproduct.toSubtype_eq_desc [DecidablePred p] : @[reassoc] theorem biproduct.ι_toSubtype_subtype (j : Subtype p) : biproduct.ι f j ≫ biproduct.toSubtype f p = biproduct.ι (Subtype.restrict p f) j := by + classical ext rw [biproduct.toSubtype, Category.assoc, biproduct.lift_π, biproduct.ι_π, biproduct.ι_π] split_ifs with h₁ h₂ h₂ @@ -819,6 +829,7 @@ section variable (f : J → C) (i : J) [HasBiproduct f] [HasBiproduct (Subtype.restrict (fun j => j ≠ i) f)] +open scoped Classical in /-- The kernel of `biproduct.π f i` is the inclusion from the biproduct which omits `i` from the index set `J` into the biproduct over `J`. -/ def biproduct.isLimitFromSubtype : @@ -844,6 +855,7 @@ instance : HasKernel (biproduct.π f i) := def kernelBiproductπIso : kernel (biproduct.π f i) ≅ ⨁ Subtype.restrict (fun j => j ≠ i) f := limit.isoLimitCone ⟨_, biproduct.isLimitFromSubtype f i⟩ +open scoped Classical in /-- The cokernel of `biproduct.ι f i` is the projection from the biproduct over the index set `J` onto the biproduct omitting `i`. -/ def biproduct.isColimitToSubtype : @@ -872,8 +884,6 @@ end section -open scoped Classical - -- Per https://github.com/leanprover-community/mathlib3/pull/15067, we only allow indexing in `Type 0` here. variable {K : Type} [Finite K] [HasFiniteBiproducts C] (f : K → C) @@ -885,6 +895,7 @@ def kernelForkBiproductToSubtype (p : Set K) : cone := KernelFork.ofι (biproduct.fromSubtype f pᶜ) (by + classical ext j k simp only [Category.assoc, biproduct.ι_fromSubtype_assoc, biproduct.ι_toSubtype_assoc, comp_zero, zero_comp] @@ -893,6 +904,7 @@ def kernelForkBiproductToSubtype (p : Set K) : isLimit := KernelFork.IsLimit.ofι _ _ (fun {_} g _ => g ≫ biproduct.toSubtype f pᶜ) (by + classical intro W' g' w ext j simp only [Category.assoc, biproduct.toSubtype_fromSubtype, Pi.compl_apply, @@ -920,6 +932,7 @@ def cokernelCoforkBiproductFromSubtype (p : Set K) : cocone := CokernelCofork.ofπ (biproduct.toSubtype f pᶜ) (by + classical ext j k simp only [Category.assoc, Pi.compl_apply, biproduct.ι_fromSubtype_assoc, biproduct.ι_toSubtype_assoc, comp_zero, zero_comp] @@ -929,6 +942,7 @@ def cokernelCoforkBiproductFromSubtype (p : Set K) : isColimit := CokernelCofork.IsColimit.ofπ _ _ (fun {_} g _ => biproduct.fromSubtype f pᶜ ≫ g) (by + classical intro W g' w ext j simp only [biproduct.toSubtype_fromSubtype_assoc, Pi.compl_apply, biproduct.ι_map_assoc] @@ -1009,11 +1023,11 @@ variable {J : Type w} variable {C : Type u} [Category.{v} C] [HasZeroMorphisms C] variable {D : Type uD} [Category.{uD'} D] [HasZeroMorphisms D] -instance biproduct.ι_mono (f : J → C) [HasBiproduct f] (b : J) : IsSplitMono (biproduct.ι f b) := - IsSplitMono.mk' { retraction := biproduct.desc <| Pi.single b (𝟙 (f b)) } +instance biproduct.ι_mono (f : J → C) [HasBiproduct f] (b : J) : IsSplitMono (biproduct.ι f b) := by + classical exact IsSplitMono.mk' { retraction := biproduct.desc <| Pi.single b (𝟙 (f b)) } -instance biproduct.π_epi (f : J → C) [HasBiproduct f] (b : J) : IsSplitEpi (biproduct.π f b) := - IsSplitEpi.mk' { section_ := biproduct.lift <| Pi.single b (𝟙 (f b)) } +instance biproduct.π_epi (f : J → C) [HasBiproduct f] (b : J) : IsSplitEpi (biproduct.π f b) := by + classical exact IsSplitEpi.mk' { section_ := biproduct.lift <| Pi.single b (𝟙 (f b)) } /-- Auxiliary lemma for `biproduct.uniqueUpToIso`. -/ theorem biproduct.conePointUniqueUpToIso_hom (f : J → C) [HasBiproduct f] {b : Bicone f} @@ -1025,6 +1039,7 @@ theorem biproduct.conePointUniqueUpToIso_hom (f : J → C) [HasBiproduct f] {b : theorem biproduct.conePointUniqueUpToIso_inv (f : J → C) [HasBiproduct f] {b : Bicone f} (hb : b.IsBilimit) : (hb.isLimit.conePointUniqueUpToIso (biproduct.isLimit _)).inv = biproduct.desc b.ι := by + classical refine biproduct.hom_ext' _ _ fun j => hb.isLimit.hom_ext fun j' => ?_ rw [Category.assoc, IsLimit.conePointUniqueUpToIso_inv_comp, Bicone.toCone_π_app, biproduct.bicone_π, biproduct.ι_desc, biproduct.ι_π, b.toCone_π_app, b.ι_π] diff --git a/Mathlib/CategoryTheory/Limits/Shapes/FiniteProducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/FiniteProducts.lean index 76f92fb28042f..21694ecd39af1 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/FiniteProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/FiniteProducts.lean @@ -17,8 +17,6 @@ universe w v u open CategoryTheory -open scoped Classical - namespace CategoryTheory.Limits variable (C : Type u) [Category.{v} C] diff --git a/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean b/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean index baf5c502afa9e..8b225b7e7ee7d 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean @@ -36,8 +36,6 @@ open CategoryTheory open CategoryTheory.Category -open scoped Classical - namespace CategoryTheory.Limits variable (C : Type u) [Category.{v} C] @@ -555,13 +553,13 @@ end Image /-- In the presence of zero morphisms, coprojections into a coproduct are (split) monomorphisms. -/ instance isSplitMono_sigma_ι {β : Type u'} [HasZeroMorphisms C] (f : β → C) - [HasColimit (Discrete.functor f)] (b : β) : IsSplitMono (Sigma.ι f b) := - IsSplitMono.mk' { retraction := Sigma.desc <| Pi.single b (𝟙 _) } + [HasColimit (Discrete.functor f)] (b : β) : IsSplitMono (Sigma.ι f b) := by + classical exact IsSplitMono.mk' { retraction := Sigma.desc <| Pi.single b (𝟙 _) } /-- In the presence of zero morphisms, projections into a product are (split) epimorphisms. -/ instance isSplitEpi_pi_π {β : Type u'} [HasZeroMorphisms C] (f : β → C) - [HasLimit (Discrete.functor f)] (b : β) : IsSplitEpi (Pi.π f b) := - IsSplitEpi.mk' { section_ := Pi.lift <| Pi.single b (𝟙 _) } + [HasLimit (Discrete.functor f)] (b : β) : IsSplitEpi (Pi.π f b) := by + classical exact IsSplitEpi.mk' { section_ := Pi.lift <| Pi.single b (𝟙 _) } /-- In the presence of zero morphisms, coprojections into a coproduct are (split) monomorphisms. -/ instance isSplitMono_coprod_inl [HasZeroMorphisms C] {X Y : C} [HasColimit (pair X Y)] : diff --git a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean index d7d052e314259..3795a5d1b7991 100644 --- a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean +++ b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean @@ -15,8 +15,6 @@ is linear in both factors. noncomputable section -open scoped Classical - namespace CategoryTheory open CategoryTheory.Limits @@ -140,6 +138,7 @@ def leftDistributor {J : Type} [Fintype J] (X : C) (f : J → C) : X ⊗ ⨁ f theorem leftDistributor_hom {J : Type} [Fintype J] (X : C) (f : J → C) : (leftDistributor X f).hom = ∑ j : J, (X ◁ biproduct.π f j) ≫ biproduct.ι (fun j => X ⊗ f j) j := by + classical ext dsimp [leftDistributor, Functor.mapBiproduct, Functor.mapBicone] erw [biproduct.lift_π] @@ -148,6 +147,7 @@ theorem leftDistributor_hom {J : Type} [Fintype J] (X : C) (f : J → C) : theorem leftDistributor_inv {J : Type} [Fintype J] (X : C) (f : J → C) : (leftDistributor X f).inv = ∑ j : J, biproduct.π _ j ≫ (X ◁ biproduct.ι f j) := by + classical ext dsimp [leftDistributor, Functor.mapBiproduct, Functor.mapBicone] simp only [Preadditive.comp_sum, biproduct.ι_π_assoc, dite_comp, zero_comp, @@ -157,28 +157,33 @@ theorem leftDistributor_inv {J : Type} [Fintype J] (X : C) (f : J → C) : @[reassoc (attr := simp)] theorem leftDistributor_hom_comp_biproduct_π {J : Type} [Fintype J] (X : C) (f : J → C) (j : J) : (leftDistributor X f).hom ≫ biproduct.π _ j = X ◁ biproduct.π _ j := by + classical simp [leftDistributor_hom, Preadditive.sum_comp, biproduct.ι_π, comp_dite] @[reassoc (attr := simp)] theorem biproduct_ι_comp_leftDistributor_hom {J : Type} [Fintype J] (X : C) (f : J → C) (j : J) : (X ◁ biproduct.ι _ j) ≫ (leftDistributor X f).hom = biproduct.ι (fun j => X ⊗ f j) j := by + classical simp [leftDistributor_hom, Preadditive.comp_sum, ← MonoidalCategory.whiskerLeft_comp_assoc, biproduct.ι_π, whiskerLeft_dite, dite_comp] @[reassoc (attr := simp)] theorem leftDistributor_inv_comp_biproduct_π {J : Type} [Fintype J] (X : C) (f : J → C) (j : J) : (leftDistributor X f).inv ≫ (X ◁ biproduct.π _ j) = biproduct.π _ j := by + classical simp [leftDistributor_inv, Preadditive.sum_comp, ← MonoidalCategory.whiskerLeft_comp, biproduct.ι_π, whiskerLeft_dite, comp_dite] @[reassoc (attr := simp)] theorem biproduct_ι_comp_leftDistributor_inv {J : Type} [Fintype J] (X : C) (f : J → C) (j : J) : biproduct.ι _ j ≫ (leftDistributor X f).inv = X ◁ biproduct.ι _ j := by + classical simp [leftDistributor_inv, Preadditive.comp_sum, ← id_tensor_comp, biproduct.ι_π_assoc, dite_comp] theorem leftDistributor_assoc {J : Type} [Fintype J] (X Y : C) (f : J → C) : (asIso (𝟙 X) ⊗ leftDistributor Y f) ≪≫ leftDistributor X _ = (α_ X Y (⨁ f)).symm ≪≫ leftDistributor (X ⊗ Y) f ≪≫ biproduct.mapIso fun _ => α_ X Y _ := by + classical ext simp only [Category.comp_id, Category.assoc, eqToHom_refl, Iso.trans_hom, Iso.symm_hom, asIso_hom, comp_zero, comp_dite, Preadditive.sum_comp, Preadditive.comp_sum, tensor_sum, @@ -196,6 +201,7 @@ def rightDistributor {J : Type} [Fintype J] (f : J → C) (X : C) : (⨁ f) ⊗ theorem rightDistributor_hom {J : Type} [Fintype J] (f : J → C) (X : C) : (rightDistributor f X).hom = ∑ j : J, (biproduct.π f j ▷ X) ≫ biproduct.ι (fun j => f j ⊗ X) j := by + classical ext dsimp [rightDistributor, Functor.mapBiproduct, Functor.mapBicone] erw [biproduct.lift_π] @@ -204,6 +210,7 @@ theorem rightDistributor_hom {J : Type} [Fintype J] (f : J → C) (X : C) : theorem rightDistributor_inv {J : Type} [Fintype J] (f : J → C) (X : C) : (rightDistributor f X).inv = ∑ j : J, biproduct.π _ j ≫ (biproduct.ι f j ▷ X) := by + classical ext dsimp [rightDistributor, Functor.mapBiproduct, Functor.mapBicone] simp only [biproduct.ι_desc, Preadditive.comp_sum, ne_eq, biproduct.ι_π_assoc, dite_comp, @@ -212,29 +219,34 @@ theorem rightDistributor_inv {J : Type} [Fintype J] (f : J → C) (X : C) : @[reassoc (attr := simp)] theorem rightDistributor_hom_comp_biproduct_π {J : Type} [Fintype J] (f : J → C) (X : C) (j : J) : (rightDistributor f X).hom ≫ biproduct.π _ j = biproduct.π _ j ▷ X := by + classical simp [rightDistributor_hom, Preadditive.sum_comp, biproduct.ι_π, comp_dite] @[reassoc (attr := simp)] theorem biproduct_ι_comp_rightDistributor_hom {J : Type} [Fintype J] (f : J → C) (X : C) (j : J) : (biproduct.ι _ j ▷ X) ≫ (rightDistributor f X).hom = biproduct.ι (fun j => f j ⊗ X) j := by + classical simp [rightDistributor_hom, Preadditive.comp_sum, ← comp_whiskerRight_assoc, biproduct.ι_π, dite_whiskerRight, dite_comp] @[reassoc (attr := simp)] theorem rightDistributor_inv_comp_biproduct_π {J : Type} [Fintype J] (f : J → C) (X : C) (j : J) : (rightDistributor f X).inv ≫ (biproduct.π _ j ▷ X) = biproduct.π _ j := by + classical simp [rightDistributor_inv, Preadditive.sum_comp, ← MonoidalCategory.comp_whiskerRight, biproduct.ι_π, dite_whiskerRight, comp_dite] @[reassoc (attr := simp)] theorem biproduct_ι_comp_rightDistributor_inv {J : Type} [Fintype J] (f : J → C) (X : C) (j : J) : biproduct.ι _ j ≫ (rightDistributor f X).inv = biproduct.ι _ j ▷ X := by + classical simp [rightDistributor_inv, Preadditive.comp_sum, ← id_tensor_comp, biproduct.ι_π_assoc, dite_comp] theorem rightDistributor_assoc {J : Type} [Fintype J] (f : J → C) (X Y : C) : (rightDistributor f X ⊗ asIso (𝟙 Y)) ≪≫ rightDistributor _ Y = α_ (⨁ f) X Y ≪≫ rightDistributor f (X ⊗ Y) ≪≫ biproduct.mapIso fun _ => (α_ _ X Y).symm := by + classical ext simp only [Category.comp_id, Category.assoc, eqToHom_refl, Iso.symm_hom, Iso.trans_hom, asIso_hom, comp_zero, comp_dite, Preadditive.sum_comp, Preadditive.comp_sum, sum_tensor, @@ -251,6 +263,7 @@ theorem leftDistributor_rightDistributor_assoc {J : Type _} [Fintype J] α_ X (⨁ f) Y ≪≫ (asIso (𝟙 X) ⊗ rightDistributor _ Y) ≪≫ leftDistributor X _ ≪≫ biproduct.mapIso fun _ => (α_ _ _ _).symm := by + classical ext simp only [Category.comp_id, Category.assoc, eqToHom_refl, Iso.symm_hom, Iso.trans_hom, asIso_hom, comp_zero, comp_dite, Preadditive.sum_comp, Preadditive.comp_sum, sum_tensor, @@ -265,6 +278,7 @@ theorem leftDistributor_rightDistributor_assoc {J : Type _} [Fintype J] @[ext] theorem leftDistributor_ext_left {J : Type} [Fintype J] {X Y : C} {f : J → C} {g h : X ⊗ ⨁ f ⟶ Y} (w : ∀ j, (X ◁ biproduct.ι f j) ≫ g = (X ◁ biproduct.ι f j) ≫ h) : g = h := by + classical apply (cancel_epi (leftDistributor X f).inv).mp ext simp? [leftDistributor_inv, Preadditive.comp_sum_assoc, biproduct.ι_π_assoc, dite_comp] says @@ -275,6 +289,7 @@ theorem leftDistributor_ext_left {J : Type} [Fintype J] {X Y : C} {f : J → C} @[ext] theorem leftDistributor_ext_right {J : Type} [Fintype J] {X Y : C} {f : J → C} {g h : X ⟶ Y ⊗ ⨁ f} (w : ∀ j, g ≫ (Y ◁ biproduct.π f j) = h ≫ (Y ◁ biproduct.π f j)) : g = h := by + classical apply (cancel_mono (leftDistributor Y f).hom).mp ext simp? [leftDistributor_hom, Preadditive.sum_comp, Preadditive.comp_sum_assoc, biproduct.ι_π, @@ -308,25 +323,25 @@ theorem leftDistributor_ext₂_right {J : Type} [Fintype J] theorem rightDistributor_ext_left {J : Type} [Fintype J] {f : J → C} {X Y : C} {g h : (⨁ f) ⊗ X ⟶ Y} (w : ∀ j, (biproduct.ι f j ▷ X) ≫ g = (biproduct.ι f j ▷ X) ≫ h) : g = h := by + classical apply (cancel_epi (rightDistributor f X).inv).mp ext simp? [rightDistributor_inv, Preadditive.comp_sum_assoc, biproduct.ι_π_assoc, dite_comp] says simp only [rightDistributor_inv, Preadditive.comp_sum_assoc, biproduct.ι_π_assoc, dite_comp, zero_comp, Finset.sum_dite_eq, Finset.mem_univ, ↓reduceIte, eqToHom_refl, Category.id_comp] - apply w @[ext] theorem rightDistributor_ext_right {J : Type} [Fintype J] {f : J → C} {X Y : C} {g h : X ⟶ (⨁ f) ⊗ Y} (w : ∀ j, g ≫ (biproduct.π f j ▷ Y) = h ≫ (biproduct.π f j ▷ Y)) : g = h := by + classical apply (cancel_mono (rightDistributor f Y).hom).mp ext simp? [rightDistributor_hom, Preadditive.sum_comp, Preadditive.comp_sum_assoc, biproduct.ι_π, comp_dite] says simp only [rightDistributor_hom, Category.assoc, Preadditive.sum_comp, biproduct.ι_π, comp_dite, comp_zero, Finset.sum_dite_eq', Finset.mem_univ, ↓reduceIte, eqToHom_refl, Category.comp_id] - apply w @[ext] diff --git a/Mathlib/CategoryTheory/Preadditive/Biproducts.lean b/Mathlib/CategoryTheory/Preadditive/Biproducts.lean index 5cbeb420311dc..49104d0c125a5 100644 --- a/Mathlib/CategoryTheory/Preadditive/Biproducts.lean +++ b/Mathlib/CategoryTheory/Preadditive/Biproducts.lean @@ -63,8 +63,6 @@ open CategoryTheory.Functor open CategoryTheory.Preadditive -open scoped Classical - universe v v' u u' noncomputable section @@ -96,6 +94,7 @@ def isBilimitOfTotal {f : J → C} (b : Bicone f) (total : ∑ j : J, b.π j ≫ erw [← Category.assoc, eq_whisker (h ⟨j⟩)] rw [reassoced] fac := fun s j => by + classical cases j simp only [sum_comp, Category.assoc, Bicone.toCone_π_app, b.ι_π, comp_dite] -- See note [dsimp, simp]. @@ -109,6 +108,7 @@ def isBilimitOfTotal {f : J → C} (b : Bicone f) (total : ∑ j : J, b.π j ≫ intro j _ erw [Category.assoc, h ⟨j⟩] fac := fun s j => by + classical cases j simp only [comp_sum, ← Category.assoc, Bicone.toCocone_ι_app, b.ι_π, dite_comp] dsimp; simp } @@ -116,6 +116,7 @@ def isBilimitOfTotal {f : J → C} (b : Bicone f) (total : ∑ j : J, b.π j ≫ theorem IsBilimit.total {f : J → C} {b : Bicone f} (i : b.IsBilimit) : ∑ j : J, b.π j ≫ b.ι j = 𝟙 b.pt := i.isLimit.hom_ext fun j => by + classical cases j simp [sum_comp, b.ι_π, comp_dite] @@ -135,6 +136,7 @@ theorem hasBiproduct_of_total {f : J → C} (b : Bicone f) def isBilimitOfIsLimit {f : J → C} (t : Bicone f) (ht : IsLimit t.toCone) : t.IsBilimit := isBilimitOfTotal _ <| ht.hom_ext fun j => by + classical cases j simp [sum_comp, t.ι_π, dite_comp, comp_dite] @@ -148,6 +150,7 @@ def biconeIsBilimitOfLimitConeOfIsLimit {f : J → C} {t : Cone (Discrete.functo def isBilimitOfIsColimit {f : J → C} (t : Bicone f) (ht : IsColimit t.toCocone) : t.IsBilimit := isBilimitOfTotal _ <| ht.hom_ext fun j => by + classical cases j simp_rw [Bicone.toCocone_ι_app, comp_sum, ← Category.assoc, t.ι_π, dite_comp] simp @@ -204,23 +207,27 @@ theorem biproduct.total : ∑ j : J, biproduct.π f j ≫ biproduct.ι f j = theorem biproduct.lift_eq {T : C} {g : ∀ j, T ⟶ f j} : biproduct.lift g = ∑ j, g j ≫ biproduct.ι f j := by + classical ext j simp only [sum_comp, biproduct.ι_π, comp_dite, biproduct.lift_π, Category.assoc, comp_zero, Finset.sum_dite_eq', Finset.mem_univ, eqToHom_refl, Category.comp_id, if_true] theorem biproduct.desc_eq {T : C} {g : ∀ j, f j ⟶ T} : biproduct.desc g = ∑ j, biproduct.π f j ≫ g j := by + classical ext j simp [comp_sum, biproduct.ι_π_assoc, dite_comp] @[reassoc] theorem biproduct.lift_desc {T U : C} {g : ∀ j, T ⟶ f j} {h : ∀ j, f j ⟶ U} : biproduct.lift g ≫ biproduct.desc h = ∑ j : J, g j ≫ h j := by + classical simp [biproduct.lift_eq, biproduct.desc_eq, comp_sum, sum_comp, biproduct.ι_π_assoc, comp_dite, dite_comp] theorem biproduct.map_eq [HasFiniteBiproducts C] {f g : J → C} {h : ∀ j, f j ⟶ g j} : biproduct.map h = ∑ j : J, biproduct.π f j ≫ h j ≫ biproduct.ι g j := by + classical ext simp [biproduct.ι_π, biproduct.ι_π_assoc, comp_sum, sum_comp, comp_dite, dite_comp] @@ -275,6 +282,7 @@ def biproduct.reindex {β γ : Type} [Finite β] (ε : β ≃ γ) · have : ε b' ≠ ε b := by simp [h] simp [biproduct.ι_π_ne _ h, biproduct.ι_π_ne _ this] inv_hom_id := by + classical cases nonempty_fintype β ext g g' by_cases h : g' = g <;> diff --git a/Mathlib/CategoryTheory/Preadditive/HomOrthogonal.lean b/Mathlib/CategoryTheory/Preadditive/HomOrthogonal.lean index 4dfeacbfed5e2..119d7d5183d77 100644 --- a/Mathlib/CategoryTheory/Preadditive/HomOrthogonal.lean +++ b/Mathlib/CategoryTheory/Preadditive/HomOrthogonal.lean @@ -36,7 +36,6 @@ This is preliminary to defining semisimple categories. -/ -open scoped Classical open Matrix CategoryTheory.Limits universe v u @@ -64,6 +63,7 @@ section variable [HasZeroMorphisms C] [HasFiniteBiproducts C] +open scoped Classical in /-- Morphisms between two direct sums over a hom orthogonal family `s : ι → C` are equivalent to block diagonal matrices, with blocks indexed by `ι`, @@ -120,6 +120,7 @@ noncomputable def matrixDecompositionAddEquiv (o : HomOrthogonal s) {α β : Typ dsimp [biproduct.components] simp } +open scoped Classical in @[simp] theorem matrixDecomposition_id (o : HomOrthogonal s) {α : Type} [Finite α] {f : α → ι} (i : ι) : o.matrixDecomposition (𝟙 (⨁ fun a => s (f a))) i = 1 := by @@ -136,6 +137,7 @@ theorem matrixDecomposition_id (o : HomOrthogonal s) {α : Type} [Finite α] {f simpa using biproduct.ι_π_ne _ (Ne.symm h) rw [this, comp_zero] +open scoped Classical in theorem matrixDecomposition_comp (o : HomOrthogonal s) {α β γ : Type} [Finite α] [Fintype β] [Finite γ] {f : α → ι} {g : β → ι} {h : γ → ι} (z : (⨁ fun a => s (f a)) ⟶ ⨁ fun b => s (g b)) (w : (⨁ fun b => s (g b)) ⟶ ⨁ fun c => s (h c)) (i : ι) : @@ -190,6 +192,7 @@ if two direct sums over `s` are isomorphic, then they have the same multipliciti theorem equiv_of_iso (o : HomOrthogonal s) {α β : Type} [Finite α] [Finite β] {f : α → ι} {g : β → ι} (i : (⨁ fun a => s (f a)) ≅ ⨁ fun b => s (g b)) : ∃ e : α ≃ β, ∀ a, g (e a) = f a := by + classical refine ⟨Equiv.ofPreimageEquiv ?_, fun a => Equiv.ofPreimageEquiv_map _ _⟩ intro c apply Nonempty.some diff --git a/Mathlib/CategoryTheory/Preadditive/Schur.lean b/Mathlib/CategoryTheory/Preadditive/Schur.lean index c5eb3daecf5ef..fd804d0e4464f 100644 --- a/Mathlib/CategoryTheory/Preadditive/Schur.lean +++ b/Mathlib/CategoryTheory/Preadditive/Schur.lean @@ -187,8 +187,7 @@ theorem finrank_hom_simple_simple_eq_zero_iff (X Y : C) [FiniteDimensional 𝕜 have := finrank_hom_simple_simple_le_one 𝕜 X Y omega -open scoped Classical - +open scoped Classical in theorem finrank_hom_simple_simple (X Y : C) [∀ X Y : C, FiniteDimensional 𝕜 (X ⟶ Y)] [Simple X] [Simple Y] : finrank 𝕜 (X ⟶ Y) = if Nonempty (X ≅ Y) then 1 else 0 := by split_ifs with h diff --git a/Mathlib/CategoryTheory/Sites/Canonical.lean b/Mathlib/CategoryTheory/Sites/Canonical.lean index 558660bca2c1d..1c9d6de317649 100644 --- a/Mathlib/CategoryTheory/Sites/Canonical.lean +++ b/Mathlib/CategoryTheory/Sites/Canonical.lean @@ -36,7 +36,6 @@ universe v u namespace CategoryTheory -open scoped Classical open CategoryTheory Category Limits Sieve variable {C : Type u} [Category.{v} C] diff --git a/Mathlib/Combinatorics/HalesJewett.lean b/Mathlib/Combinatorics/HalesJewett.lean index 32e369960eca8..0f4879e749f4f 100644 --- a/Mathlib/Combinatorics/HalesJewett.lean +++ b/Mathlib/Combinatorics/HalesJewett.lean @@ -65,7 +65,6 @@ combinatorial line, Ramsey theory, arithmetic progression -/ open Function -open scoped Classical universe u v variable {η α ι κ : Type*} @@ -107,6 +106,7 @@ lemma coe_apply (l : Subspace η α ι) (x : η → α) (i : ι) : l x i = (l.id -- Note: This is not made a `FunLike` instance to avoid having two syntactically different coercions lemma coe_injective [Nontrivial α] : Injective ((⇑) : Subspace η α ι → (η → α) → ι → α) := by + classical rintro l m hlm ext i simp only [funext_iff] at hlm @@ -453,6 +453,7 @@ end Line monoid, and `S` is a finite subset, then there exists a monochromatic homothetic copy of `S`. -/ theorem exists_mono_homothetic_copy {M κ : Type*} [AddCommMonoid M] (S : Finset M) [Finite κ] (C : M → κ) : ∃ a > 0, ∃ (b : M) (c : κ), ∀ s ∈ S, C (a • s + b) = c := by + classical obtain ⟨ι, _inst, hι⟩ := Line.exists_mono_in_high_dimension S κ specialize hι fun v => C <| ∑ i, v i obtain ⟨l, c, hl⟩ := hι diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean index 6d9c7efcbf156..f6d6c39558155 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean @@ -110,8 +110,6 @@ theorem not_isUniform_iff : unfold IsUniform simp only [not_forall, not_lt, exists_prop, exists_and_left, Rat.cast_abs, Rat.cast_sub] -open scoped Classical - variable (G) /-- An arbitrary pair of subsets witnessing the non-uniformity of `(s, t)`. If `(s, t)` is uniform, @@ -149,6 +147,7 @@ theorem nonuniformWitnesses_spec (h : ¬G.IsUniform ε s t) : rw [nonuniformWitnesses, dif_pos h] exact (not_isUniform_iff.1 h).choose_spec.2.choose_spec.2.2.2 +open scoped Classical in /-- Arbitrary witness of non-uniformity. `G.nonuniformWitness ε s t` and `G.nonuniformWitness ε t s` form a pair of subsets witnessing the non-uniformity of `(s, t)`. If `(s, t)` is uniform, returns `s`. -/ diff --git a/Mathlib/Data/Finite/Card.lean b/Mathlib/Data/Finite/Card.lean index a1193eb881ff5..c47971873d694 100644 --- a/Mathlib/Data/Finite/Card.lean +++ b/Mathlib/Data/Finite/Card.lean @@ -28,8 +28,6 @@ it. We generally put such theorems into the `SetTheory.Cardinal.Finite` module. noncomputable section -open scoped Classical - variable {α β γ : Type*} /-- There is (noncomputably) an equivalence between a finite type `α` and `Fin (Nat.card α)`. -/ @@ -42,6 +40,7 @@ def Finite.equivFinOfCardEq [Finite α] {n : ℕ} (h : Nat.card α = n) : α ≃ subst h apply Finite.equivFin +open scoped Classical in theorem Nat.card_eq (α : Type*) : Nat.card α = if _ : Finite α then @Fintype.card α (Fintype.ofFinite α) else 0 := by cases finite_or_infinite α @@ -93,6 +92,7 @@ theorem card_le_of_embedding [Finite β] (f : α ↪ β) : Nat.card α ≤ Nat.c theorem card_le_of_surjective [Finite α] (f : α → β) (hf : Function.Surjective f) : Nat.card β ≤ Nat.card α := by + classical haveI := Fintype.ofFinite α haveI := Fintype.ofSurjective f hf simpa only [Nat.card_eq_fintype_card] using Fintype.card_le_of_surjective f hf @@ -152,11 +152,13 @@ theorem card_range_le [Finite α] (f : α → β) : Nat.card (Set.range f) ≤ N card_le_of_surjective _ Set.surjective_onto_range theorem card_subtype_le [Finite α] (p : α → Prop) : Nat.card { x // p x } ≤ Nat.card α := by + classical haveI := Fintype.ofFinite α simpa only [Nat.card_eq_fintype_card] using Fintype.card_subtype_le p theorem card_subtype_lt [Finite α] {p : α → Prop} {x : α} (hx : ¬p x) : Nat.card { x // p x } < Nat.card α := by + classical haveI := Fintype.ofFinite α simpa only [Nat.card_eq_fintype_card, gt_iff_lt] using Fintype.card_subtype_lt hx diff --git a/Mathlib/Data/Finite/Prod.lean b/Mathlib/Data/Finite/Prod.lean index c0dbebd034aaa..9d577868e5da9 100644 --- a/Mathlib/Data/Finite/Prod.lean +++ b/Mathlib/Data/Finite/Prod.lean @@ -14,8 +14,6 @@ import Mathlib.Data.Fintype.Vector assert_not_exists OrderedRing assert_not_exists MonoidWithZero -open scoped Classical - variable {α β : Type*} namespace Finite @@ -38,6 +36,7 @@ end Finite instance Pi.finite {α : Sort*} {β : α → Sort*} [Finite α] [∀ a, Finite (β a)] : Finite (∀ a, β a) := by + classical haveI := Fintype.ofFinite (PLift α) haveI := fun a => Fintype.ofFinite (PLift (β a)) exact @@ -45,6 +44,7 @@ instance Pi.finite {α : Sort*} {β : α → Sort*} [Finite α] [∀ a, Finite ( (Equiv.piCongr Equiv.plift fun _ => Equiv.plift) instance [Finite α] {n : ℕ} : Finite (Sym α n) := by + classical haveI := Fintype.ofFinite α infer_instance @@ -117,8 +117,6 @@ Some set instances do not appear here since they are consequences of others, for namespace Finite.Set -open scoped Classical - instance finite_prod (s : Set α) (t : Set β) [Finite s] [Finite t] : Finite (s ×ˢ t : Set (α × β)) := Finite.of_equiv _ (Equiv.Set.prod s t).symm diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 2a6a0a1559942..b6901aef2d20c 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -438,8 +438,6 @@ theorem subset_union_elim {s : Finset α} {t₁ t₂ : Set α} (h : ↑s ⊆ t section Classical -open scoped Classical - -- Porting note: The notation `{ x ∈ s | p x }` in Lean 4 is hardcoded to be about `Set`. -- So at the moment the whole `Sep`-class is useless, as it doesn't have notation. -- /-- The following instance allows us to write `{x ∈ s | p x}` for `Finset.filter p s`. diff --git a/Mathlib/Data/Finset/Finsupp.lean b/Mathlib/Data/Finset/Finsupp.lean index 43e04648e3ccf..670a4f1ca9509 100644 --- a/Mathlib/Data/Finset/Finsupp.lean +++ b/Mathlib/Data/Finset/Finsupp.lean @@ -31,19 +31,20 @@ noncomputable section open Finsupp -open scoped Classical open Pointwise variable {ι α : Type*} [Zero α] {s : Finset ι} {f : ι →₀ α} namespace Finset +open scoped Classical in /-- Finitely supported product of finsets. -/ protected def finsupp (s : Finset ι) (t : ι → Finset α) : Finset (ι →₀ α) := (s.pi t).map ⟨indicator s, indicator_injective s⟩ theorem mem_finsupp_iff {t : ι → Finset α} : f ∈ s.finsupp t ↔ f.support ⊆ s ∧ ∀ i ∈ s, f i ∈ t i := by + classical refine mem_map.trans ⟨?_, ?_⟩ · rintro ⟨f, hf, rfl⟩ refine ⟨support_indicator_subset _ _, fun i hi => ?_⟩ @@ -70,8 +71,8 @@ theorem mem_finsupp_iff_of_support_subset {t : ι →₀ Finset α} (ht : t.supp · rwa [H, mem_zero] at h @[simp] -theorem card_finsupp (s : Finset ι) (t : ι → Finset α) : #(s.finsupp t) = ∏ i ∈ s, #(t i) := - (card_map _).trans <| card_pi _ _ +theorem card_finsupp (s : Finset ι) (t : ι → Finset α) : #(s.finsupp t) = ∏ i ∈ s, #(t i) := by + classical exact (card_map _).trans <| card_pi _ _ end Finset diff --git a/Mathlib/Data/Finset/Functor.lean b/Mathlib/Data/Finset/Functor.lean index 3d5642f536c25..7656a4984033b 100644 --- a/Mathlib/Data/Finset/Functor.lean +++ b/Mathlib/Data/Finset/Functor.lean @@ -191,18 +191,19 @@ theorem id_traverse [DecidableEq α] (s : Finset α) : traverse (pure : α → I rw [traverse, Multiset.id_traverse] exact s.val_toFinset -open scoped Classical - +open scoped Classical in @[simp] theorem map_comp_coe (h : α → β) : Functor.map h ∘ Multiset.toFinset = Multiset.toFinset ∘ Functor.map h := funext fun _ => image_toFinset +open scoped Classical in @[simp] theorem map_comp_coe_apply (h : α → β) (s : Multiset α) : s.toFinset.image h = (h <$> s).toFinset := congrFun (map_comp_coe h) s +open scoped Classical in theorem map_traverse (g : α → G β) (h : β → γ) (s : Finset α) : Functor.map h <$> traverse g s = traverse (Functor.map h ∘ g) s := by unfold traverse diff --git a/Mathlib/Data/Finsupp/Interval.lean b/Mathlib/Data/Finsupp/Interval.lean index 9fed894fefb53..de42113086356 100644 --- a/Mathlib/Data/Finsupp/Interval.lean +++ b/Mathlib/Data/Finsupp/Interval.lean @@ -25,10 +25,7 @@ supported. noncomputable section -open Finset Finsupp Function - -open scoped Classical -open Pointwise +open Finset Finsupp Function Pointwise variable {ι α : Type*} @@ -56,6 +53,7 @@ section RangeIcc variable [Zero α] [PartialOrder α] [LocallyFiniteOrder α] {f g : ι →₀ α} {i : ι} {a : α} +open scoped Classical in /-- Pointwise `Finset.Icc` bundled as a `Finsupp`. -/ @[simps toFun] def rangeIcc (f g : ι →₀ α) : ι →₀ Finset α where @@ -71,6 +69,7 @@ def rangeIcc (f g : ι →₀ α) : ι →₀ Finset α where -- Porting note: Added as alternative to rangeIcc_toFun to be used in proof of card_Icc lemma coe_rangeIcc (f g : ι →₀ α) : rangeIcc f g i = Icc (f i) (g i) := rfl +open scoped Classical in @[simp] theorem rangeIcc_support (f g : ι →₀ α) : (rangeIcc f g).support = f.support ∪ g.support := rfl @@ -83,6 +82,7 @@ section PartialOrder variable [PartialOrder α] [Zero α] [LocallyFiniteOrder α] (f g : ι →₀ α) +open scoped Classical in instance instLocallyFiniteOrder : LocallyFiniteOrder (ι →₀ α) := -- Porting note: Not needed (due to open scoped Classical), in mathlib3 too -- haveI := Classical.decEq ι @@ -94,20 +94,25 @@ instance instLocallyFiniteOrder : LocallyFiniteOrder (ι →₀ α) := simp_rw [mem_rangeIcc_apply_iff] exact forall_and +open scoped Classical in theorem Icc_eq : Icc f g = (f.support ∪ g.support).finsupp (f.rangeIcc g) := rfl +open scoped Classical in -- Porting note: removed [DecidableEq ι] theorem card_Icc : #(Icc f g) = ∏ i ∈ f.support ∪ g.support, #(Icc (f i) (g i)):= by simp_rw [Icc_eq, card_finsupp, coe_rangeIcc] +open scoped Classical in -- Porting note: removed [DecidableEq ι] theorem card_Ico : #(Ico f g) = ∏ i ∈ f.support ∪ g.support, #(Icc (f i) (g i)) - 1 := by rw [card_Ico_eq_card_Icc_sub_one, card_Icc] +open scoped Classical in -- Porting note: removed [DecidableEq ι] theorem card_Ioc : #(Ioc f g) = ∏ i ∈ f.support ∪ g.support, #(Icc (f i) (g i)) - 1 := by rw [card_Ioc_eq_card_Icc_sub_one, card_Icc] +open scoped Classical in -- Porting note: removed [DecidableEq ι] theorem card_Ioo : #(Ioo f g) = ∏ i ∈ f.support ∪ g.support, #(Icc (f i) (g i)) - 2 := by rw [card_Ioo_eq_card_Icc_sub_two, card_Icc] @@ -117,6 +122,7 @@ end PartialOrder section Lattice variable [Lattice α] [Zero α] [LocallyFiniteOrder α] (f g : ι →₀ α) +open scoped Classical in -- Porting note: removed [DecidableEq ι] theorem card_uIcc : #(uIcc f g) = ∏ i ∈ f.support ∪ g.support, #(uIcc (f i) (g i)) := by diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index 5075621a3a24a..08568f59f0f7d 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -871,8 +871,7 @@ noncomputable def fintypeOfNotInfinite {α : Type*} (h : ¬Infinite α) : Fintyp section -open scoped Classical - +open scoped Classical in /-- Any type is (classically) either a `Fintype`, or `Infinite`. One can obtain the relevant typeclasses via `cases fintypeOrInfinite α`. diff --git a/Mathlib/Data/Fintype/Order.lean b/Mathlib/Data/Fintype/Order.lean index 78516f9aea242..67390d4754196 100644 --- a/Mathlib/Data/Fintype/Order.lean +++ b/Mathlib/Data/Fintype/Order.lean @@ -80,8 +80,7 @@ section BoundedOrder variable (α) -open scoped Classical - +open scoped Classical in -- See note [reducible non-instances] /-- A finite bounded lattice is complete. -/ noncomputable abbrev toCompleteLattice [Lattice α] [BoundedOrder α] : CompleteLattice α where diff --git a/Mathlib/Data/Fintype/Prod.lean b/Mathlib/Data/Fintype/Prod.lean index 882ddcb288b7e..804d624e70fc9 100644 --- a/Mathlib/Data/Fintype/Prod.lean +++ b/Mathlib/Data/Fintype/Prod.lean @@ -55,8 +55,6 @@ theorem Fintype.card_prod (α β : Type*) [Fintype α] [Fintype β] : section -open scoped Classical - @[simp] theorem infinite_prod : Infinite (α × β) ↔ Infinite α ∧ Nonempty β ∨ Nonempty α ∧ Infinite β := by refine @@ -69,6 +67,7 @@ theorem infinite_prod : Infinite (α × β) ↔ Infinite α ∧ Nonempty β ∨ instance Pi.infinite_of_left {ι : Sort*} {π : ι → Type*} [∀ i, Nontrivial <| π i] [Infinite ι] : Infinite (∀ i : ι, π i) := by + classical choose m n hm using fun i => exists_pair_ne (π i) refine Infinite.of_injective (fun i => update m i (n i)) fun x y h => of_not_not fun hne => ?_ simp_rw [update_eq_iff, update_of_ne hne] at h @@ -76,9 +75,10 @@ instance Pi.infinite_of_left {ι : Sort*} {π : ι → Type*} [∀ i, Nontrivial /-- If at least one `π i` is infinite and the rest nonempty, the pi type of all `π` is infinite. -/ theorem Pi.infinite_of_exists_right {ι : Sort*} {π : ι → Sort*} (i : ι) [Infinite <| π i] - [∀ i, Nonempty <| π i] : Infinite (∀ i : ι, π i) := + [∀ i, Nonempty <| π i] : Infinite (∀ i : ι, π i) := by + classical let ⟨m⟩ := @Pi.instNonempty ι π _ - Infinite.of_injective _ (update_injective m i) + exact Infinite.of_injective _ (update_injective m i) /-- See `Pi.infinite_of_exists_right` for the case that only one `π i` is infinite. -/ instance Pi.infinite_of_right {ι : Sort*} {π : ι → Type*} [∀ i, Infinite <| π i] [Nonempty ι] : diff --git a/Mathlib/Data/Fintype/Sum.lean b/Mathlib/Data/Fintype/Sum.lean index 930f4d0abf5d9..7ed85d9190de4 100644 --- a/Mathlib/Data/Fintype/Sum.lean +++ b/Mathlib/Data/Fintype/Sum.lean @@ -122,8 +122,6 @@ theorem Fintype.card_subtype_or_disjoint (p q : α → Prop) (h : Disjoint p q) section -open scoped Classical - @[simp] theorem infinite_sum : Infinite (α ⊕ β) ↔ Infinite α ∨ Infinite β := by refine ⟨fun H => ?_, fun H => H.elim (@Sum.infinite_of_left α β) (@Sum.infinite_of_right α β)⟩ diff --git a/Mathlib/Data/Real/Archimedean.lean b/Mathlib/Data/Real/Archimedean.lean index 1fc5a7ad21d35..039acaddf3405 100644 --- a/Mathlib/Data/Real/Archimedean.lean +++ b/Mathlib/Data/Real/Archimedean.lean @@ -16,7 +16,6 @@ import Mathlib.Order.Interval.Set.Disjoint assert_not_exists Finset -open scoped Classical open Pointwise CauSeq namespace Real @@ -111,9 +110,11 @@ theorem exists_isGLB (hne : s.Nonempty) (hbdd : BddBelow s) : ∃ x, IsGLB s x : rw [← isLUB_neg] exact Classical.choose_spec (Real.exists_isLUB hne' hbdd') +open scoped Classical in noncomputable instance : SupSet ℝ := ⟨fun s => if h : s.Nonempty ∧ BddAbove s then Classical.choose (exists_isLUB h.1 h.2) else 0⟩ +open scoped Classical in theorem sSup_def (s : Set ℝ) : sSup s = if h : s.Nonempty ∧ BddAbove s then Classical.choose (exists_isLUB h.1 h.2) else 0 := rfl @@ -253,6 +254,7 @@ lemma iInf_nonneg (hf : ∀ i, 0 ≤ f i) : 0 ≤ iInf f := Real.le_iInf hf le_r /-- As `sSup s = 0` when `s` is a set of reals that's unbounded above, it suffices to show that `s` contains a nonnegative element to show that `0 ≤ sSup s`. -/ lemma sSup_nonneg' (hs : ∃ x ∈ s, 0 ≤ x) : 0 ≤ sSup s := by + classical obtain ⟨x, hxs, hx⟩ := hs exact dite _ (fun h ↦ le_csSup_of_le h hxs hx) fun h ↦ (sSup_of_not_bddAbove h).ge @@ -263,6 +265,7 @@ lemma iSup_nonneg' (hf : ∃ i, 0 ≤ f i) : 0 ≤ ⨆ i, f i := sSup_nonneg' <| /-- As `sInf s = 0` when `s` is a set of reals that's unbounded below, it suffices to show that `s` contains a nonpositive element to show that `sInf s ≤ 0`. -/ lemma sInf_nonpos' (hs : ∃ x ∈ s, x ≤ 0) : sInf s ≤ 0 := by + classical obtain ⟨x, hxs, hx⟩ := hs exact dite _ (fun h ↦ csInf_le_of_le h hxs hx) fun h ↦ (sInf_of_not_bddBelow h).le diff --git a/Mathlib/Data/Real/Basic.lean b/Mathlib/Data/Real/Basic.lean index a47daa561e74e..b16bdaec3d484 100644 --- a/Mathlib/Data/Real/Basic.lean +++ b/Mathlib/Data/Real/Basic.lean @@ -483,8 +483,6 @@ instance : SemilatticeInf ℝ := instance : SemilatticeSup ℝ := inferInstance -open scoped Classical - instance leTotal_R : IsTotal ℝ (· ≤ ·) := ⟨by intros a b @@ -492,6 +490,7 @@ instance leTotal_R : IsTotal ℝ (· ≤ ·) := induction' b using Real.ind_mk with b simpa using le_total a b⟩ +open scoped Classical in noncomputable instance linearOrder : LinearOrder ℝ := Lattice.toLinearOrder ℝ diff --git a/Mathlib/Data/Real/ENatENNReal.lean b/Mathlib/Data/Real/ENatENNReal.lean index 2e04581738900..559f57bd15842 100644 --- a/Mathlib/Data/Real/ENatENNReal.lean +++ b/Mathlib/Data/Real/ENatENNReal.lean @@ -14,7 +14,6 @@ In this file we define a coercion from `ℕ∞` to `ℝ≥0∞` and prove some b assert_not_exists Finset -open scoped Classical open NNReal ENNReal noncomputable section diff --git a/Mathlib/Data/Real/Hyperreal.lean b/Mathlib/Data/Real/Hyperreal.lean index 18ffdc0c6a65a..fd849459ea9f9 100644 --- a/Mathlib/Data/Real/Hyperreal.lean +++ b/Mathlib/Data/Real/Hyperreal.lean @@ -11,7 +11,6 @@ import Mathlib.Analysis.SpecificLimits.Basic -/ -open scoped Classical open Filter Germ Topology /-- Hyperreal numbers on the ultrafilter extending the cofinite filter -/ @@ -181,6 +180,7 @@ theorem epsilon_lt_pos (x : ℝ) : 0 < x → ε < x := def IsSt (x : ℝ*) (r : ℝ) := ∀ δ : ℝ, 0 < δ → (r - δ : ℝ*) < x ∧ x < r + δ +open scoped Classical in /-- Standard part function: like a "round" to ℝ instead of ℤ -/ noncomputable def st : ℝ* → ℝ := fun x => if h : ∃ r, IsSt x r then Classical.choose h else 0 @@ -525,10 +525,11 @@ theorem not_infinite_mul {x y : ℝ*} (hx : ¬Infinite x) (hy : ¬Infinite y) : theorem st_add {x y : ℝ*} (hx : ¬Infinite x) (hy : ¬Infinite y) : st (x + y) = st x + st y := (isSt_st' (not_infinite_add hx hy)).unique ((isSt_st' hx).add (isSt_st' hy)) -theorem st_neg (x : ℝ*) : st (-x) = -st x := - if h : Infinite x then by - rw [h.st_eq, (infinite_neg.2 h).st_eq, neg_zero] - else (isSt_st' (not_infinite_neg h)).unique (isSt_st' h).neg +theorem st_neg (x : ℝ*) : st (-x) = -st x := by + classical + by_cases h : Infinite x + · rw [h.st_eq, (infinite_neg.2 h).st_eq, neg_zero] + · exact (isSt_st' (not_infinite_neg h)).unique (isSt_st' h).neg theorem st_mul {x y : ℝ*} (hx : ¬Infinite x) (hy : ¬Infinite y) : st (x * y) = st x * st y := have hx' := isSt_st' hx diff --git a/Mathlib/Data/Set/BoolIndicator.lean b/Mathlib/Data/Set/BoolIndicator.lean index 09d93329c0505..95d7f96641fe5 100644 --- a/Mathlib/Data/Set/BoolIndicator.lean +++ b/Mathlib/Data/Set/BoolIndicator.lean @@ -35,8 +35,7 @@ theorem preimage_boolIndicator_true : s.boolIndicator ⁻¹' {true} = s := theorem preimage_boolIndicator_false : s.boolIndicator ⁻¹' {false} = sᶜ := ext fun x ↦ (s.not_mem_iff_boolIndicator x).symm -open scoped Classical - +open scoped Classical in theorem preimage_boolIndicator_eq_union (t : Set Bool) : s.boolIndicator ⁻¹' t = (if true ∈ t then s else ∅) ∪ if false ∈ t then sᶜ else ∅ := by ext x diff --git a/Mathlib/Data/Set/Countable.lean b/Mathlib/Data/Set/Countable.lean index b22bd0f331322..8b90a104ce95a 100644 --- a/Mathlib/Data/Set/Countable.lean +++ b/Mathlib/Data/Set/Countable.lean @@ -24,7 +24,6 @@ sets, countable set noncomputable section -open scoped Classical open Function Set Encodable universe u v w x diff --git a/Mathlib/Data/Set/Finite/Basic.lean b/Mathlib/Data/Set/Finite/Basic.lean index 0ee4e5ce237b5..c775fbbcad06a 100644 --- a/Mathlib/Data/Set/Finite/Basic.lean +++ b/Mathlib/Data/Set/Finite/Basic.lean @@ -405,8 +405,6 @@ Some set instances do not appear here since they are consequences of others, for namespace Finite.Set -open scoped Classical - example {s : Set α} [Finite α] : Finite s := inferInstance @@ -419,10 +417,12 @@ example (a : α) : Finite ({a} : Set α) := instance finite_union (s t : Set α) [Finite s] [Finite t] : Finite (s ∪ t : Set α) := by cases nonempty_fintype s cases nonempty_fintype t + classical infer_instance instance finite_sep (s : Set α) (p : α → Prop) [Finite s] : Finite ({ a ∈ s | p a } : Set α) := by cases nonempty_fintype s + classical infer_instance protected theorem subset (s : Set α) {t : Set α} [Finite s] (h : t ⊆ s) : Finite t := by @@ -443,6 +443,7 @@ instance finite_insert (a : α) (s : Set α) [Finite s] : Finite (insert a s : S instance finite_image (s : Set α) (f : α → β) [Finite s] : Finite (f '' s) := by cases nonempty_fintype s + classical infer_instance end Finite.Set diff --git a/Mathlib/Data/Set/Finite/Lattice.lean b/Mathlib/Data/Set/Finite/Lattice.lean index d48b9b82b810e..85c9d8dbf4beb 100644 --- a/Mathlib/Data/Set/Finite/Lattice.lean +++ b/Mathlib/Data/Set/Finite/Lattice.lean @@ -72,8 +72,6 @@ Some set instances do not appear here since they are consequences of others, for namespace Finite.Set -open scoped Classical - instance finite_iUnion [Finite ι] (f : ι → Set α) [∀ i, Finite (f i)] : Finite (⋃ i, f i) := by rw [iUnion_eq_range_psigma] apply Set.finite_range diff --git a/Mathlib/Data/Set/Finite/Monad.lean b/Mathlib/Data/Set/Finite/Monad.lean index 4614681a6df64..eb54e4a5396ba 100644 --- a/Mathlib/Data/Set/Finite/Monad.lean +++ b/Mathlib/Data/Set/Finite/Monad.lean @@ -78,8 +78,6 @@ Some set instances do not appear here since they are consequences of others, for namespace Finite.Set -open scoped Classical - theorem finite_pure (a : α) : (pure a : Set α).Finite := toFinite _ diff --git a/Mathlib/Data/Set/Finite/Range.lean b/Mathlib/Data/Set/Finite/Range.lean index fdbe1c806503c..d1f22cc150285 100644 --- a/Mathlib/Data/Set/Finite/Range.lean +++ b/Mathlib/Data/Set/Finite/Range.lean @@ -58,9 +58,8 @@ Some set instances do not appear here since they are consequences of others, for namespace Finite.Set -open scoped Classical - instance finite_range (f : ι → α) [Finite ι] : Finite (range f) := by + classical haveI := Fintype.ofFinite (PLift ι) infer_instance diff --git a/Mathlib/Dynamics/PeriodicPts.lean b/Mathlib/Dynamics/PeriodicPts.lean index 3c2e4112481df..28ed4478ac2af 100644 --- a/Mathlib/Dynamics/PeriodicPts.lean +++ b/Mathlib/Dynamics/PeriodicPts.lean @@ -228,16 +228,16 @@ variable {f} theorem Semiconj.mapsTo_periodicPts {g : α → β} (h : Semiconj g fa fb) : MapsTo g (periodicPts fa) (periodicPts fb) := fun _ ⟨n, hn, hx⟩ => ⟨n, hn, hx.map h⟩ -open scoped Classical - noncomputable section +open scoped Classical in /-- Minimal period of a point `x` under an endomorphism `f`. If `x` is not a periodic point of `f`, then `minimalPeriod f x = 0`. -/ def minimalPeriod (f : α → α) (x : α) := if h : x ∈ periodicPts f then Nat.find h else 0 theorem isPeriodicPt_minimalPeriod (f : α → α) (x : α) : IsPeriodicPt f (minimalPeriod f x) x := by + classical delta minimalPeriod split_ifs with hx · exact (Nat.find_spec hx).2 @@ -258,6 +258,7 @@ theorem iterate_mod_minimalPeriod_eq : f^[n % minimalPeriod f x] x = f^[n] x := (isPeriodicPt_minimalPeriod f x).iterate_mod_apply n theorem minimalPeriod_pos_of_mem_periodicPts (hx : x ∈ periodicPts f) : 0 < minimalPeriod f x := by + classical simp only [minimalPeriod, dif_pos hx, (Nat.find_spec hx).1.lt] theorem minimalPeriod_eq_zero_of_nmem_periodicPts (hx : x ∉ periodicPts f) : @@ -276,6 +277,7 @@ theorem minimalPeriod_eq_zero_iff_nmem_periodicPts : minimalPeriod f x = 0 ↔ x theorem IsPeriodicPt.minimalPeriod_le (hn : 0 < n) (hx : IsPeriodicPt f n x) : minimalPeriod f x ≤ n := by + classical rw [minimalPeriod, dif_pos (mk_mem_periodicPts hn hx)] exact Nat.find_min' (mk_mem_periodicPts hn hx) ⟨hn, hx⟩ diff --git a/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean b/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean index 6f5ada837fa50..80d8d3e437a38 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean @@ -32,13 +32,14 @@ namespace NumberField open Module NumberField NumberField.InfinitePlace Matrix -open scoped Classical Real nonZeroDivisors +open scoped Real nonZeroDivisors variable (K : Type*) [Field K] [NumberField K] open MeasureTheory MeasureTheory.Measure ZSpan NumberField.mixedEmbedding NumberField.InfinitePlace ENNReal NNReal Complex +open scoped Classical in theorem _root_.NumberField.mixedEmbedding.volume_fundamentalDomain_latticeBasis : volume (fundamentalDomain (latticeBasis K)) = (2 : ℝ≥0∞)⁻¹ ^ nrComplexPlaces K * sqrt ‖discr K‖₊ := by @@ -73,6 +74,7 @@ theorem _root_.NumberField.mixedEmbedding.volume_fundamentalDomain_latticeBasis stdBasis_repr_eq_matrixToStdBasis_mul K _ (fun _ => rfl)] rfl +open scoped Classical in theorem _root_.NumberField.mixedEmbedding.covolume_integerLattice : ZLattice.covolume (mixedEmbedding.integerLattice K) = (2 ⁻¹) ^ nrComplexPlaces K * √|discr K| := by @@ -81,6 +83,7 @@ theorem _root_.NumberField.mixedEmbedding.covolume_integerLattice : ENNReal.toReal_inv, toReal_ofNat, ENNReal.coe_toReal, Real.coe_sqrt, coe_nnnorm, Int.norm_eq_abs] +open scoped Classical in theorem _root_.NumberField.mixedEmbedding.covolume_idealLattice (I : (FractionalIdeal (𝓞 K)⁰ K)ˣ) : ZLattice.covolume (mixedEmbedding.idealLattice K I) = (FractionalIdeal.absNorm (I : FractionalIdeal (𝓞 K)⁰ K)) * @@ -95,6 +98,7 @@ theorem exists_ne_zero_mem_ideal_of_norm_le_mul_sqrt_discr (I : (FractionalIdeal ∃ a ∈ (I : FractionalIdeal (𝓞 K)⁰ K), a ≠ 0 ∧ |Algebra.norm ℚ (a : K)| ≤ FractionalIdeal.absNorm I.1 * (4 / π) ^ nrComplexPlaces K * (finrank ℚ K).factorial / (finrank ℚ K) ^ (finrank ℚ K) * Real.sqrt |discr K| := by + classical -- The smallest possible value for `exists_ne_zero_mem_ideal_of_norm_le` let B := (minkowskiBound K I * (convexBodySumFactor K)⁻¹).toReal ^ (1 / (finrank ℚ K : ℝ)) have h_le : (minkowskiBound K I) ≤ volume (convexBodySum K B) := by @@ -315,6 +319,7 @@ theorem finite_of_discr_bdd_of_isReal : {K : { F : IntermediateField ℚ A // FiniteDimensional ℚ F} | haveI : NumberField K := @NumberField.mk _ _ inferInstance K.prop {w : InfinitePlace K | IsReal w}.Nonempty ∧ |discr K| ≤ N }.Finite := by + classical -- The bound on the degree of the generating polynomials let D := rankOfDiscrBdd N -- The bound on the Minkowski bound @@ -362,6 +367,7 @@ theorem finite_of_discr_bdd_of_isComplex : {K : { F : IntermediateField ℚ A // FiniteDimensional ℚ F} | haveI : NumberField K := @NumberField.mk _ _ inferInstance K.prop {w : InfinitePlace K | IsComplex w}.Nonempty ∧ |discr K| ≤ N }.Finite := by + classical -- The bound on the degree of the generating polynomials let D := rankOfDiscrBdd N -- The bound on the Minkowski bound diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index 4fbfe52301fe9..1d17123834ee7 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -58,7 +58,6 @@ spans the full space over `ℝ` (see `unitLattice_span_eq_top`); this is the mai see the section `span_top` below for more details. -/ -open scoped Classical open Finset variable {K} @@ -85,6 +84,7 @@ variable {K} theorem logEmbedding_component (x : (𝓞 K)ˣ) (w : {w : InfinitePlace K // w ≠ w₀}) : (logEmbedding K (Additive.ofMul x)) w = mult w.val * Real.log (w.val x) := rfl +open scoped Classical in theorem sum_logEmbedding_component (x : (𝓞 K)ˣ) : ∑ w, logEmbedding K (Additive.ofMul x) w = - mult (w₀ : InfinitePlace K) * Real.log (w₀ (x : K)) := by @@ -125,12 +125,14 @@ theorem logEmbedding_eq_zero_iff {x : (𝓞 K)ˣ} : · ext w rw [logEmbedding_component, h w.val, Real.log_one, mul_zero, Pi.zero_apply] +open scoped Classical in theorem logEmbedding_component_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : ‖logEmbedding K x‖ ≤ r) (w : {w : InfinitePlace K // w ≠ w₀}) : |logEmbedding K (Additive.ofMul x) w| ≤ r := by lift r to NNReal using hr simp_rw [Pi.norm_def, NNReal.coe_le_coe, Finset.sup_le_iff, ← NNReal.coe_le_coe] at h exact h w (mem_univ _) +open scoped Classical in theorem log_le_of_logEmbedding_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : ‖logEmbedding K (Additive.ofMul x)‖ ≤ r) (w : InfinitePlace K) : |Real.log (w x)| ≤ (Fintype.card (InfinitePlace K)) * r := by @@ -164,6 +166,7 @@ noncomputable def _root_.NumberField.Units.unitLattice : Submodule ℤ ({w : InfinitePlace K // w ≠ w₀} → ℝ) := Submodule.map (logEmbedding K).toIntLinearMap ⊤ +open scoped Classical in theorem unitLattice_inter_ball_finite (r : ℝ) : ((unitLattice K : Set ({ w : InfinitePlace K // w ≠ w₀} → ℝ)) ∩ Metric.closedBall 0 r).Finite := by @@ -305,6 +308,7 @@ theorem exists_unit (w₁ : InfinitePlace K) : theorem unitLattice_span_eq_top : Submodule.span ℝ (unitLattice K : Set ({w : InfinitePlace K // w ≠ w₀} → ℝ)) = ⊤ := by + classical refine le_antisymm le_top ?_ -- The standard basis let B := Pi.basisFun ℝ {w : InfinitePlace K // w ≠ w₀} @@ -340,13 +344,13 @@ section statements variable [NumberField K] -open scoped Classical open dirichletUnitTheorem Module /-- The unit rank of the number field `K`, it is equal to `card (InfinitePlace K) - 1`. -/ def rank : ℕ := Fintype.card (InfinitePlace K) - 1 instance instDiscrete_unitLattice : DiscreteTopology (unitLattice K) := by + classical refine discreteTopology_of_isOpen_singleton_zero ?_ refine isOpen_singleton_of_finite_mem_nhds 0 (s := Metric.closedBall 0 1) ?_ ?_ · exact Metric.closedBall_mem_nhds _ (by norm_num) @@ -357,17 +361,20 @@ instance instDiscrete_unitLattice : DiscreteTopology (unitLattice K) := by rintro ⟨x, hx, rfl⟩ exact ⟨Subtype.mem x, hx⟩ +open scoped Classical in instance instZLattice_unitLattice : IsZLattice ℝ (unitLattice K) where span_top := unitLattice_span_eq_top K protected theorem finrank_eq_rank : finrank ℝ ({w : InfinitePlace K // w ≠ w₀} → ℝ) = Units.rank K := by + classical simp only [finrank_fintype_fun_eq_card, Fintype.card_subtype_compl, Fintype.card_ofSubsingleton, rank] @[simp] theorem unitLattice_rank : finrank ℤ (unitLattice K) = Units.rank K := by + classical rw [← Units.finrank_eq_rank, ZLattice.rank ℝ] /-- The map obtained by quotienting by the kernel of `logEmbedding`. -/ @@ -422,11 +429,11 @@ theorem logEmbeddingEquiv_apply (x : (𝓞 K)ˣ) : logEmbeddingEquiv K (Additive.ofMul (QuotientGroup.mk x)) = logEmbedding K (Additive.ofMul x) := rfl -instance : Module.Free ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := - Module.Free.of_equiv (logEmbeddingEquiv K).symm +instance : Module.Free ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := by + classical exact Module.Free.of_equiv (logEmbeddingEquiv K).symm -instance : Module.Finite ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := - Module.Finite.equiv (logEmbeddingEquiv K).symm +instance : Module.Finite ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := by + classical exact Module.Finite.equiv (logEmbeddingEquiv K).symm -- Note that we prove this instance first and then deduce from it the instance -- `Monoid.FG (𝓞 K)ˣ`, and not the other way around, due to no `Subgroup` version diff --git a/Mathlib/RingTheory/AlgebraicIndependent/RankAndCardinality.lean b/Mathlib/RingTheory/AlgebraicIndependent/RankAndCardinality.lean index 1c736c384c401..e5df26045cf9a 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent/RankAndCardinality.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent/RankAndCardinality.lean @@ -31,8 +31,6 @@ noncomputable section open Function Set Subalgebra MvPolynomial Algebra -open scoped Classical - universe u v w open AlgebraicIndependent diff --git a/Mathlib/RingTheory/AlgebraicIndependent/TranscendenceBasis.lean b/Mathlib/RingTheory/AlgebraicIndependent/TranscendenceBasis.lean index db169e8b20d4c..9c1d64bd63250 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent/TranscendenceBasis.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent/TranscendenceBasis.lean @@ -32,8 +32,6 @@ noncomputable section open Function Set Subalgebra MvPolynomial Algebra -open scoped Classical - universe u v w variable {ι ι' : Type*} (R : Type*) {K A A' : Type*} diff --git a/Mathlib/RingTheory/AlgebraicIndependent/Transcendental.lean b/Mathlib/RingTheory/AlgebraicIndependent/Transcendental.lean index f49b84def844f..e6c9a512fb501 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent/Transcendental.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent/Transcendental.lean @@ -25,8 +25,6 @@ noncomputable section open Function Set Subalgebra MvPolynomial Algebra -open scoped Classical - variable {ι R A : Type*} {x : ι → A} variable [CommRing R] [CommRing A] [Algebra R A] diff --git a/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean b/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean index 0aac9bcf5b559..058f95c8bd4ad 100644 --- a/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean +++ b/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean @@ -58,7 +58,7 @@ dedekind domain, dedekind ring, adic valuation noncomputable section -open scoped Classical Multiplicative +open scoped Multiplicative open Multiplicative IsDedekindDomain @@ -69,7 +69,7 @@ namespace IsDedekindDomain.HeightOneSpectrum /-! ### Adic valuations on the Dedekind domain R -/ - +open scoped Classical in /-- The additive `v`-adic valuation of `r ∈ R` is the exponent of `v` in the factorization of the ideal `(r)`, if `r` is nonzero, or infinity, if `r = 0`. `intValuationDef` is the corresponding multiplicative valuation. -/ @@ -87,6 +87,7 @@ theorem intValuationDef_if_pos {r : R} (hr : r = 0) : v.intValuationDef r = 0 := theorem intValuationDef_zero : v.intValuationDef 0 = 0 := if_pos rfl +open scoped Classical in theorem intValuationDef_if_neg {r : R} (hr : r ≠ 0) : v.intValuationDef r = Multiplicative.ofAdd @@ -131,6 +132,7 @@ alias int_valuation_le_one := intValuation_le_one /-- The `v`-adic valuation of `r ∈ R` is less than 1 if and only if `v` divides the ideal `(r)`. -/ theorem intValuation_lt_one_iff_dvd (r : R) : v.intValuationDef r < 1 ↔ v.asIdeal ∣ Ideal.span {r} := by + classical rw [intValuationDef] split_ifs with hr · simp [hr] @@ -148,6 +150,7 @@ alias int_valuation_lt_one_iff_dvd := intValuation_lt_one_iff_dvd `vⁿ` divides the ideal `(r)`. -/ theorem intValuation_le_pow_iff_dvd (r : R) (n : ℕ) : v.intValuationDef r ≤ Multiplicative.ofAdd (-(n : ℤ)) ↔ v.asIdeal ^ n ∣ Ideal.span {r} := by + classical rw [intValuationDef] split_ifs with hr · simp_rw [hr, Ideal.dvd_span_singleton, zero_le', Submodule.zero_mem] @@ -168,6 +171,7 @@ alias IntValuation.map_zero' := intValuation.map_zero' /-- The `v`-adic valuation of `1 : R` equals 1. -/ theorem intValuation.map_one' : v.intValuationDef 1 = 1 := by + classical rw [v.intValuationDef_if_neg (zero_ne_one.symm : (1 : R) ≠ 0), Ideal.span_singleton_one, ← Ideal.one_eq_top, Associates.mk_one, Associates.factors_one, Associates.count_zero (by apply v.associates_irreducible), Int.ofNat_zero, neg_zero, ofAdd_zero, @@ -179,6 +183,7 @@ alias IntValuation.map_one' := intValuation.map_one' /-- The `v`-adic valuation of a product equals the product of the valuations. -/ theorem intValuation.map_mul' (x y : R) : v.intValuationDef (x * y) = v.intValuationDef x * v.intValuationDef y := by + classical simp only [intValuationDef] by_cases hx : x = 0 · rw [hx, zero_mul, if_pos (Eq.refl _), zero_mul] @@ -206,6 +211,7 @@ alias IntValuation.le_max_iff_min_le := intValuation.le_max_iff_min_le /-- The `v`-adic valuation of a sum is bounded above by the maximum of the valuations. -/ theorem intValuation.map_add_le_max' (x y : R) : v.intValuationDef (x + y) ≤ max (v.intValuationDef x) (v.intValuationDef y) := by + classical by_cases hx : x = 0 · rw [hx, zero_add] conv_rhs => rw [intValuationDef, if_pos (Eq.refl _)] @@ -257,6 +263,7 @@ theorem intValuation_apply {r : R} (v : IsDedekindDomain.HeightOneSpectrum R) : /-- There exists `π ∈ R` with `v`-adic valuation `Multiplicative.ofAdd (-1)`. -/ theorem intValuation_exists_uniformizer : ∃ π : R, v.intValuationDef π = Multiplicative.ofAdd (-1 : ℤ) := by + classical have hv : _root_.Irreducible (Associates.mk v.asIdeal) := v.associates_irreducible have hlt : v.asIdeal ^ 2 < v.asIdeal := by rw [← Ideal.dvdNotUnit_iff_lt] @@ -284,6 +291,7 @@ alias int_valuation_exists_uniformizer := intValuation_exists_uniformizer /-- The `I`-adic valuation of a generator of `I` equals `(-1 : ℤₘ₀)` -/ theorem intValuation_singleton {r : R} (hr : r ≠ 0) (hv : v.asIdeal = Ideal.span {r}) : v.intValuation r = Multiplicative.ofAdd (-1 : ℤ) := by + classical rw [intValuation_apply, v.intValuationDef_if_neg hr, ← hv, Associates.count_self, Int.ofNat_one, ofAdd_neg, WithZero.coe_inv] apply v.associates_irreducible diff --git a/Mathlib/RingTheory/DedekindDomain/Different.lean b/Mathlib/RingTheory/DedekindDomain/Different.lean index be9c17f20d9c5..7b195346a4cf3 100644 --- a/Mathlib/RingTheory/DedekindDomain/Different.lean +++ b/Mathlib/RingTheory/DedekindDomain/Different.lean @@ -200,12 +200,11 @@ lemma isIntegral_discr_mul_of_mem_traceDual variable (A K) -open scoped Classical - variable [IsDomain A] [IsFractionRing B L] [Nontrivial B] [NoZeroDivisors B] namespace FractionalIdeal +open scoped Classical in /-- The dual of a non-zero fractional ideal is the dual of the submodule under the traceform. -/ noncomputable def dual (I : FractionalIdeal B⁰ L) : @@ -238,8 +237,6 @@ variable [IsFractionRing B L] [IsIntegrallyClosed A] open Submodule -open scoped Classical - local notation:max I:max "ᵛ" => Submodule.traceDual A K I variable [IsDedekindDomain B] {I J : FractionalIdeal B⁰ L} @@ -322,8 +319,9 @@ lemma le_dual_iff (hJ : J ≠ 0) : 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)) + I⁻¹ ≤ dual A K I := by + classical + exact 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 diff --git a/Mathlib/RingTheory/DedekindDomain/Factorization.lean b/Mathlib/RingTheory/DedekindDomain/Factorization.lean index 61803b29f13c4..85a552bcab99e 100644 --- a/Mathlib/RingTheory/DedekindDomain/Factorization.lean +++ b/Mathlib/RingTheory/DedekindDomain/Factorization.lean @@ -47,7 +47,7 @@ dedekind domain, fractional ideal, ideal, factorization noncomputable section -open scoped Classical nonZeroDivisors +open scoped nonZeroDivisors open Set Function UniqueFactorizationMonoid IsDedekindDomain IsDedekindDomain.HeightOneSpectrum Classical diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 21c7cf5ee7db4..fe0bf42cd4973 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -1252,8 +1252,7 @@ noncomputable def IsDedekindDomain.quotientEquivPiOfProdEq {ι : Type*} [Fintype (hPp.le_of_pow_le hPi)).trans <| Eq.symm <| (Ring.DimensionLeOne.prime_le_prime_iff_eq (prime j).ne_zero).mp (hPp.le_of_pow_le hPj) -open scoped Classical - +open scoped Classical in /-- **Chinese remainder theorem** for a Dedekind domain: `R ⧸ I` factors as `Π i, R ⧸ (P i ^ e i)`, where `P i` ranges over the prime factors of `I` and `e i` over the multiplicities. -/ noncomputable def IsDedekindDomain.quotientEquivPiFactors {I : Ideal R} (hI : I ≠ ⊥) : diff --git a/Mathlib/RingTheory/Flat/FaithfullyFlat.lean b/Mathlib/RingTheory/Flat/FaithfullyFlat.lean index 203fd8a1cd50f..39e1cd3c439b7 100644 --- a/Mathlib/RingTheory/Flat/FaithfullyFlat.lean +++ b/Mathlib/RingTheory/Flat/FaithfullyFlat.lean @@ -188,11 +188,10 @@ lemma of_linearEquiv {N : Type*} [AddCommGroup N] [Module R N] [FaithfullyFlat R section -open Classical - /-- A direct sum of faithfully flat `R`-modules is faithfully flat. -/ instance directSum {ι : Type*} [Nonempty ι] (M : ι → Type*) [∀ i, AddCommGroup (M i)] [∀ i, Module R (M i)] [∀ i, FaithfullyFlat R (M i)] : FaithfullyFlat R (⨁ i, M i) := by + classical rw [iff_flat_and_lTensor_faithful] refine ⟨inferInstance, fun N _ _ hN ↦ ?_⟩ obtain ⟨i⟩ := ‹Nonempty ι› @@ -202,8 +201,8 @@ instance directSum {ι : Type*} [Nonempty ι] (M : ι → Type*) [∀ i, AddComm apply (TensorProduct.directSumLeft R M N).toEquiv.nontrivial /-- Free `R`-modules over discrete types are flat. -/ -instance finsupp (ι : Type v) [Nonempty ι] : FaithfullyFlat R (ι →₀ R) := - of_linearEquiv _ _ (finsuppLEquivDirectSum R R ι) +instance finsupp (ι : Type v) [Nonempty ι] : FaithfullyFlat R (ι →₀ R) := by + classical exact of_linearEquiv _ _ (finsuppLEquivDirectSum R R ι) end diff --git a/Mathlib/RingTheory/HahnSeries/Addition.lean b/Mathlib/RingTheory/HahnSeries/Addition.lean index 0a73b0ddc2ff1..817b97d65f6d7 100644 --- a/Mathlib/RingTheory/HahnSeries/Addition.lean +++ b/Mathlib/RingTheory/HahnSeries/Addition.lean @@ -24,8 +24,6 @@ coefficients in `R`, whose supports are partially well-ordered. With further str open Finset Function -open scoped Classical - noncomputable section variable {Γ Γ' R S U V : Type*} @@ -109,6 +107,7 @@ lemma addOppositeEquiv_symm_support (x : (HahnSeries Γ R)ᵃᵒᵖ) : @[simp] lemma addOppositeEquiv_orderTop (x : HahnSeries Γ (Rᵃᵒᵖ)) : (addOppositeEquiv x).unop.orderTop = x.orderTop := by + classical simp only [orderTop, AddOpposite.unop_op, mk_eq_zero, EmbeddingLike.map_eq_zero_iff, addOppositeEquiv_support, ne_eq] simp only [addOppositeEquiv_apply, AddOpposite.unop_op, mk_eq_zero, zero_coeff] @@ -123,6 +122,7 @@ lemma addOppositeEquiv_symm_orderTop (x : (HahnSeries Γ R)ᵃᵒᵖ) : @[simp] lemma addOppositeEquiv_leadingCoeff (x : HahnSeries Γ (Rᵃᵒᵖ)) : (addOppositeEquiv x).unop.leadingCoeff = x.leadingCoeff.unop := by + classical simp only [leadingCoeff, AddOpposite.unop_op, mk_eq_zero, EmbeddingLike.map_eq_zero_iff, addOppositeEquiv_support, ne_eq] simp only [addOppositeEquiv_apply, AddOpposite.unop_op, mk_eq_zero, zero_coeff] @@ -276,10 +276,11 @@ protected lemma map_neg [AddGroup S] (f : R →+ S) {x : HahnSeries Γ R} : ext; simp theorem orderTop_neg {x : HahnSeries Γ R} : (-x).orderTop = x.orderTop := by - simp only [orderTop, support_neg, neg_eq_zero] + classical simp only [orderTop, support_neg, neg_eq_zero] @[simp] theorem order_neg [Zero Γ] {f : HahnSeries Γ R} : (-f).order = f.order := by + classical by_cases hf : f = 0 · simp only [hf, neg_zero] simp only [order, support_neg, neg_eq_zero] diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean index 04233a83a2559..5b810da90a937 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean @@ -18,7 +18,6 @@ We prove basic properties of `IsIntegralClosure`. -/ -open scoped Classical open Polynomial Submodule section inv diff --git a/Mathlib/RingTheory/LocalProperties/Basic.lean b/Mathlib/RingTheory/LocalProperties/Basic.lean index 354847e24b104..8e5d4900156be 100644 --- a/Mathlib/RingTheory/LocalProperties/Basic.lean +++ b/Mathlib/RingTheory/LocalProperties/Basic.lean @@ -41,7 +41,7 @@ In this file, we define local properties in general. -/ -open scoped Pointwise Classical +open scoped Pointwise universe u diff --git a/Mathlib/RingTheory/Localization/Integral.lean b/Mathlib/RingTheory/Localization/Integral.lean index cb310c7e3e585..5e9059d15096e 100644 --- a/Mathlib/RingTheory/Localization/Integral.lean +++ b/Mathlib/RingTheory/Localization/Integral.lean @@ -37,8 +37,7 @@ open Polynomial variable [IsLocalization M S] -open scoped Classical - +open scoped Classical in /-- `coeffIntegerNormalization p` gives the coefficients of the polynomial `integerNormalization p` -/ noncomputable def coeffIntegerNormalization (p : S[X]) (i : ℕ) : R := @@ -71,6 +70,7 @@ theorem integerNormalization_coeff (p : S[X]) (i : ℕ) : theorem integerNormalization_spec (p : S[X]) : ∃ b : M, ∀ i, algebraMap R S ((integerNormalization M p).coeff i) = (b : R) • p.coeff i := by + classical use Classical.choose (exist_integer_multiples_of_finset M (p.support.image p.coeff)) intro i rw [integerNormalization_coeff, coeffIntegerNormalization] diff --git a/Mathlib/RingTheory/Polynomial/IrreducibleRing.lean b/Mathlib/RingTheory/Polynomial/IrreducibleRing.lean index 914f39fd47f38..0796dcb8e167d 100644 --- a/Mathlib/RingTheory/Polynomial/IrreducibleRing.lean +++ b/Mathlib/RingTheory/Polynomial/IrreducibleRing.lean @@ -25,8 +25,6 @@ polynomial, irreducible ring, nilradical, prime ideal -/ -open scoped Classical Polynomial - open Polynomial noncomputable section diff --git a/Mathlib/RingTheory/Polynomial/SeparableDegree.lean b/Mathlib/RingTheory/Polynomial/SeparableDegree.lean index aa17885393628..247b2449134af 100644 --- a/Mathlib/RingTheory/Polynomial/SeparableDegree.lean +++ b/Mathlib/RingTheory/Polynomial/SeparableDegree.lean @@ -37,7 +37,6 @@ noncomputable section namespace Polynomial -open scoped Classical open Polynomial section CommSemiring diff --git a/Mathlib/RingTheory/PowerSeries/Inverse.lean b/Mathlib/RingTheory/PowerSeries/Inverse.lean index aecb07876b39f..5d54bce963615 100644 --- a/Mathlib/RingTheory/PowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/PowerSeries/Inverse.lean @@ -231,8 +231,7 @@ theorem Inv_divided_by_X_pow_order_leftInv {f : k⟦X⟧} (hf : f ≠ 0) : rw [mul_comm] exact mul_invOfUnit (divided_by_X_pow_order hf) (firstUnitCoeff hf) rfl -open scoped Classical - +open scoped Classical in /-- `Unit_of_divided_by_X_pow_order` is the unit power series obtained by dividing a non-zero power series by the largest power of `X` that divides it. -/ def Unit_of_divided_by_X_pow_order (f : k⟦X⟧) : k⟦X⟧ˣ := @@ -352,8 +351,9 @@ theorem normUnit_X : normUnit (X : k⟦X⟧) = 1 := by theorem X_eq_normalizeX : (X : k⟦X⟧) = normalize X := by simp only [normalize_apply, normUnit_X, Units.val_one, mul_one] -open UniqueFactorizationMonoid Classical +open UniqueFactorizationMonoid +open scoped Classical in theorem normalized_count_X_eq_of_coe {P : k[X]} (hP : P ≠ 0) : Multiset.count PowerSeries.X (normalizedFactors (P : k⟦X⟧)) = Multiset.count Polynomial.X (normalizedFactors P) := by diff --git a/Mathlib/RingTheory/RingHom/Finite.lean b/Mathlib/RingTheory/RingHom/Finite.lean index 338b996f159e9..f42863038ca92 100644 --- a/Mathlib/RingTheory/RingHom/Finite.lean +++ b/Mathlib/RingTheory/RingHom/Finite.lean @@ -52,7 +52,7 @@ theorem finite_isStableUnderBaseChange : IsStableUnderBaseChange @Finite := by end RingHom -open scoped Pointwise Classical +open scoped Pointwise universe u @@ -114,6 +114,7 @@ theorem RingHom.localization_away_map_finite (r : R) [IsLocalization.Away r R'] [IsLocalization.Away (f r) S'] (hf : f.Finite) : (IsLocalization.Away.map R' S' f r).Finite := finite_localizationPreserves.away f r _ _ hf +open scoped Classical in /-- Let `S` be an `R`-algebra, `M` a submonoid of `R`, and `S' = M⁻¹S`. If the image of some `x : S` falls in the span of some finite `s ⊆ S'` over `R`, then there exists some `m : M` such that `m • x` falls in the @@ -194,6 +195,7 @@ theorem multiple_mem_adjoin_of_mem_localization_adjoin [Algebra R' S] [Algebra R /-- `S` is a finite `R`-algebra if there exists a set `{ r }` that spans `R` such that `Sᵣ` is a finite `Rᵣ`-algebra. -/ theorem RingHom.finite_ofLocalizationSpan : RingHom.OfLocalizationSpan @RingHom.Finite := by + classical rw [RingHom.ofLocalizationSpan_iff_finite] introv R hs H -- We first setup the instances diff --git a/Mathlib/RingTheory/RingHom/FiniteType.lean b/Mathlib/RingTheory/RingHom/FiniteType.lean index 592be263a4219..6defe23560662 100644 --- a/Mathlib/RingTheory/RingHom/FiniteType.lean +++ b/Mathlib/RingTheory/RingHom/FiniteType.lean @@ -26,7 +26,7 @@ Let `R` be a commutative ring, `S` is an `R`-algebra, `M` be a submonoid of `R`. namespace RingHom -open scoped Pointwise TensorProduct Classical +open scoped Pointwise TensorProduct universe u @@ -40,6 +40,7 @@ theorem finiteType_stableUnderComposition : StableUnderComposition @FiniteType : /-- If `S` is a finite type `R`-algebra, then `S' = M⁻¹S` is a finite type `R' = M⁻¹R`-algebra. -/ theorem finiteType_localizationPreserves : RingHom.LocalizationPreserves @RingHom.FiniteType := by + classical introv R hf -- mirrors the proof of `localization_map_finite` letI := f.toAlgebra @@ -75,6 +76,7 @@ theorem localization_away_map_finiteType (r : R) [IsLocalization.Away r R'] variable {S'} +open scoped Classical in /-- Let `S` be an `R`-algebra, `M` a submonoid of `S`, `S' = M⁻¹S`. Suppose the image of some `x : S` falls in the adjoin of some finite `s ⊆ S'` over `R`, and `A` is an `R`-subalgebra of `S` containing both `M` and the numerators of `s`. @@ -100,6 +102,7 @@ theorem IsLocalization.exists_smul_mem_of_mem_adjoin [Algebra R S] [Algebra R S' rw [Submonoid.smul_def, smul_eq_mul, Submonoid.coe_mul, SubmonoidClass.coe_pow, mul_assoc, ← ha₂, mul_comm] +open scoped Classical in /-- Let `S` be an `R`-algebra, `M` a submonoid of `R`, and `S' = M⁻¹S`. If the image of some `x : S` falls in the adjoin of some finite `s ⊆ S'` over `R`, then there exists some `m : M` such that `m • x` falls in the @@ -118,6 +121,7 @@ theorem IsLocalization.lift_mem_adjoin_finsetIntegerMultiple [Algebra R S] [Alge simpa only [Submonoid.smul_def, algebraMap_smul] using e theorem finiteType_ofLocalizationSpan : RingHom.OfLocalizationSpan @RingHom.FiniteType := by + classical rw [RingHom.ofLocalizationSpan_iff_finite] introv R hs H -- mirrors the proof of `finite_ofLocalizationSpan` diff --git a/Mathlib/RingTheory/RootsOfUnity/Basic.lean b/Mathlib/RingTheory/RootsOfUnity/Basic.lean index a4a2e5f1d6ed4..9025712363bb4 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Basic.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Basic.lean @@ -34,8 +34,6 @@ Note that `rootsOfUnity 0 M` is the top subgroup of `Mˣ` (as the condition `ζ^ satisfied for all units). -/ -open scoped Classical - noncomputable section open Polynomial @@ -197,13 +195,15 @@ theorem rootsOfUnityEquivNthRoots_symm_apply (x : { x // x ∈ nthRoots k (1 : R variable (k R) -instance rootsOfUnity.fintype : Fintype (rootsOfUnity k R) := - Fintype.ofEquiv { x // x ∈ nthRoots k (1 : R) } (rootsOfUnityEquivNthRoots R k).symm +instance rootsOfUnity.fintype : Fintype (rootsOfUnity k R) := by + classical + exact Fintype.ofEquiv { x // x ∈ nthRoots k (1 : R) } (rootsOfUnityEquivNthRoots R k).symm instance rootsOfUnity.isCyclic : IsCyclic (rootsOfUnity k R) := isCyclic_of_subgroup_isDomain ((Units.coeHom R).comp (rootsOfUnity k R).subtype) coe_injective -theorem card_rootsOfUnity : Fintype.card (rootsOfUnity k R) ≤ k := +theorem card_rootsOfUnity : Fintype.card (rootsOfUnity k R) ≤ k := by + classical calc Fintype.card (rootsOfUnity k R) = Fintype.card { x // x ∈ nthRoots k (1 : R) } := Fintype.card_congr (rootsOfUnityEquivNthRoots R k) diff --git a/Mathlib/RingTheory/RootsOfUnity/PrimitiveRoots.lean b/Mathlib/RingTheory/RootsOfUnity/PrimitiveRoots.lean index f1f8616f7cb18..731cc72273b2b 100644 --- a/Mathlib/RingTheory/RootsOfUnity/PrimitiveRoots.lean +++ b/Mathlib/RingTheory/RootsOfUnity/PrimitiveRoots.lean @@ -40,8 +40,6 @@ of the `Units`), but lemmas like `IsPrimitiveRoot.isUnit` and `IsPrimitiveRoot.coe_units_iff` should provide the necessary glue. -/ -open scoped Classical - noncomputable section open Polynomial Finset @@ -66,6 +64,7 @@ section primitiveRoots variable {k : ℕ} +open scoped Classical in /-- `primitiveRoots k R` is the finset of primitive `k`-th roots of unity in the integral domain `R`. -/ def primitiveRoots (k : ℕ) (R : Type*) [CommRing R] [IsDomain R] : Finset R := @@ -76,11 +75,13 @@ variable [CommRing R] [IsDomain R] -- TODO?: replace `(h0 : 0 < k)` by `[NeZero k]` @[simp] theorem mem_primitiveRoots {ζ : R} (h0 : 0 < k) : ζ ∈ primitiveRoots k R ↔ IsPrimitiveRoot ζ k := by + classical rw [primitiveRoots, mem_filter, Multiset.mem_toFinset, mem_nthRoots h0, and_iff_right_iff_imp] exact IsPrimitiveRoot.pow_eq_one @[simp] theorem primitiveRoots_zero : primitiveRoots 0 R = ∅ := by + classical rw [primitiveRoots, nthRoots_zero, Multiset.toFinset_zero, Finset.filter_empty] theorem isPrimitiveRoot_of_mem_primitiveRoots {ζ : R} (h : ζ ∈ primitiveRoots k R) : @@ -566,6 +567,7 @@ theorem nthRoots_eq {n : ℕ} {ζ : R} (hζ : IsPrimitiveRoot ζ n) {α a : R} ( pow_mul, hζ.pow_eq_one, one_pow, one_mul] · simpa only [Multiset.card_map, Multiset.card_range] using card_nthRoots n a +open scoped Classical in theorem card_nthRoots {n : ℕ} {ζ : R} (hζ : IsPrimitiveRoot ζ n) (a : R) : Multiset.card (nthRoots n a) = if ∃ α, α ^ n = a then n else 0 := by split_ifs with h @@ -620,6 +622,7 @@ theorem nthRoots_one_nodup {ζ : R} {n : ℕ} (h : IsPrimitiveRoot ζ n) : @[simp] theorem card_nthRootsFinset {ζ : R} {n : ℕ} (h : IsPrimitiveRoot ζ n) : #(nthRootsFinset n R) = n := by + classical rw [nthRootsFinset, ← Multiset.toFinset_eq (nthRoots_one_nodup h), card_mk, h.card_nthRoots_one] open scoped Nat @@ -651,6 +654,7 @@ theorem disjoint {k l : ℕ} (h : k ≠ l) : Disjoint (primitiveRoots k R) (prim h <| (isPrimitiveRoot_of_mem_primitiveRoots hk).unique <| isPrimitiveRoot_of_mem_primitiveRoots hl +open scoped Classical in /-- `nthRoots n` as a `Finset` is equal to the union of `primitiveRoots i R` for `i ∣ n` if there is a primitive `n`th root of unity in `R`. -/ private -- marking as `private` since `nthRoots_one_eq_biUnion_primitiveRoots` can be used instead @@ -678,6 +682,7 @@ theorem nthRoots_one_eq_biUnion_primitiveRoots' {ζ : R} {n : ℕ} [NeZero n] · intro i _ j _ hdiff exact disjoint hdiff +open scoped Classical in /-- `nthRoots n` as a `Finset` is equal to the union of `primitiveRoots i R` for `i ∣ n` if there is a primitive `n`th root of unity in `R`. -/ theorem nthRoots_one_eq_biUnion_primitiveRoots {ζ : R} {n : ℕ} diff --git a/Mathlib/RingTheory/Smooth/StandardSmooth.lean b/Mathlib/RingTheory/Smooth/StandardSmooth.lean index aafe269b9c441..9aef55d125b25 100644 --- a/Mathlib/RingTheory/Smooth/StandardSmooth.lean +++ b/Mathlib/RingTheory/Smooth/StandardSmooth.lean @@ -83,7 +83,7 @@ in June 2024. universe t t' w w' u v -open TensorProduct MvPolynomial Classical +open TensorProduct MvPolynomial variable (n m : ℕ) @@ -180,6 +180,7 @@ instance (h : Function.Bijective (algebraMap R S)) : Fintype (ofBijectiveAlgebra @[simp] lemma ofBijectiveAlgebraMap_jacobian (h : Function.Bijective (algebraMap R S)) : (ofBijectiveAlgebraMap h).jacobian = 1 := by + classical have : (algebraMap (ofBijectiveAlgebraMap h).Ring S).mapMatrix (ofBijectiveAlgebraMap h).jacobiMatrix = 1 := by ext (i j : PEmpty) @@ -263,6 +264,7 @@ the lower-right block has determinant jacobian of `P`. variable [Fintype (Q.comp P).rels] +open scoped Classical in private lemma jacobiMatrix_comp_inl_inr (i : Q.rels) (j : P.rels) : (Q.comp P).jacobiMatrix (Sum.inl i) (Sum.inr j) = 0 := by rw [jacobiMatrix_apply] @@ -270,6 +272,7 @@ private lemma jacobiMatrix_comp_inl_inr (i : Q.rels) (j : P.rels) : apply MvPolynomial.vars_rename at hmem simp at hmem +open scoped Classical in private lemma jacobiMatrix_comp_₁₂ : (Q.comp P).jacobiMatrix.toBlocks₁₂ = 0 := by ext i j : 1 simp [Matrix.toBlocks₁₂, jacobiMatrix_comp_inl_inr] @@ -278,6 +281,7 @@ section Q variable [Fintype Q.rels] +open scoped Classical in private lemma jacobiMatrix_comp_inl_inl (i j : Q.rels) : aeval (Sum.elim X (MvPolynomial.C ∘ P.val)) ((Q.comp P).jacobiMatrix (Sum.inl j) (Sum.inl i)) = Q.jacobiMatrix j i := by @@ -285,6 +289,7 @@ private lemma jacobiMatrix_comp_inl_inl (i j : Q.rels) : ← Q.comp_aeval_relation_inl P.toPresentation] apply aeval_sum_elim_pderiv_inl +open scoped Classical in private lemma jacobiMatrix_comp_₁₁_det : (aeval (Q.comp P).val) (Q.comp P).jacobiMatrix.toBlocks₁₁.det = Q.jacobian := by rw [jacobian_eq_jacobiMatrix_det, AlgHom.map_det (aeval (Q.comp P).val), RingHom.map_det] @@ -299,6 +304,7 @@ section P variable [Fintype P.rels] +open scoped Classical in private lemma jacobiMatrix_comp_inr_inr (i j : P.rels) : (Q.comp P).jacobiMatrix (Sum.inr i) (Sum.inr j) = MvPolynomial.rename Sum.inr (P.jacobiMatrix i j) := by @@ -306,6 +312,7 @@ private lemma jacobiMatrix_comp_inr_inr (i j : P.rels) : simp only [comp_map, Sum.elim_inr] apply pderiv_rename Sum.inr_injective +open scoped Classical in private lemma jacobiMatrix_comp_₂₂_det : (aeval (Q.comp P).val) (Q.comp P).jacobiMatrix.toBlocks₂₂.det = algebraMap S T P.jacobian := by rw [jacobian_eq_jacobiMatrix_det] @@ -332,6 +339,7 @@ end /-- The jacobian of the composition of presentations is the product of the jacobians. -/ @[simp] lemma comp_jacobian_eq_jacobian_smul_jacobian : (Q.comp P).jacobian = P.jacobian • Q.jacobian := by + classical cases nonempty_fintype Q.rels cases nonempty_fintype P.rels letI : Fintype (Q.comp P).rels := inferInstanceAs <| Fintype (Q.rels ⊕ P.rels) diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean index 5f137806b6aa5..a907a17caf885 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean @@ -299,12 +299,11 @@ end UniqueFactorizationMonoid namespace UniqueFactorizationMonoid -open scoped Classical - open Multiset Associates variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] +open scoped Classical in /-- Noncomputably defines a `normalizationMonoid` structure on a `UniqueFactorizationMonoid`. -/ protected noncomputable def normalizationMonoid : NormalizationMonoid α := normalizationMonoidOfMonoidHomRightInverse diff --git a/Mathlib/RingTheory/Valuation/ValuationSubring.lean b/Mathlib/RingTheory/Valuation/ValuationSubring.lean index 8f4342e738c9a..8a87a29540687 100644 --- a/Mathlib/RingTheory/Valuation/ValuationSubring.lean +++ b/Mathlib/RingTheory/Valuation/ValuationSubring.lean @@ -23,8 +23,6 @@ The order structure on `ValuationSubring K`. universe u -open scoped Classical - noncomputable section variable (K : Type u) [Field K] @@ -349,10 +347,12 @@ def primeSpectrumOrderEquiv : (PrimeSpectrum A)ᵒᵈ ≃o {S // A ≤ S} := all_goals exact le_ofPrime A (PrimeSpectrum.asIdeal _), fun h => by apply ofPrime_le_of_le; exact h⟩ } -instance le_total_ideal : IsTotal {S // A ≤ S} LE.le := +instance le_total_ideal : IsTotal {S // A ≤ S} LE.le := by + classical let _ : IsTotal (PrimeSpectrum A) (· ≤ ·) := ⟨fun ⟨x, _⟩ ⟨y, _⟩ => LE.isTotal.total x y⟩ - ⟨(primeSpectrumOrderEquiv A).symm.toRelEmbedding.isTotal.total⟩ + exact ⟨(primeSpectrumOrderEquiv A).symm.toRelEmbedding.isTotal.total⟩ +open scoped Classical in instance linearOrderOverring : LinearOrder {S // A ≤ S} where le_total := (le_total_ideal A).1 max_def a b := congr_fun₂ sup_eq_maxDefault a b diff --git a/Mathlib/RingTheory/WittVector/Domain.lean b/Mathlib/RingTheory/WittVector/Domain.lean index e3d68f08bd302..f3c3235ed4e54 100644 --- a/Mathlib/RingTheory/WittVector/Domain.lean +++ b/Mathlib/RingTheory/WittVector/Domain.lean @@ -34,8 +34,6 @@ the 0th component of which must be nonzero. noncomputable section -open scoped Classical - namespace WittVector open Function @@ -81,6 +79,7 @@ theorem eq_iterate_verschiebung {x : 𝕎 R} {n : ℕ} (h : ∀ i < n, x.coeff i theorem verschiebung_nonzero {x : 𝕎 R} (hx : x ≠ 0) : ∃ n : ℕ, ∃ x' : 𝕎 R, x'.coeff 0 ≠ 0 ∧ x = verschiebung^[n] x' := by + classical have hex : ∃ k : ℕ, x.coeff k ≠ 0 := by by_contra! hall apply hx diff --git a/Mathlib/RingTheory/WittVector/InitTail.lean b/Mathlib/RingTheory/WittVector/InitTail.lean index 0568831e3df4f..79b87d594512e 100644 --- a/Mathlib/RingTheory/WittVector/InitTail.lean +++ b/Mathlib/RingTheory/WittVector/InitTail.lean @@ -44,12 +44,11 @@ namespace WittVector open MvPolynomial -open scoped Classical - noncomputable section section +open scoped Classical in /-- `WittVector.select P x`, for a predicate `P : ℕ → Prop` is the Witt vector whose `n`-th coefficient is `x.coeff n` if `P n` is true, and `0` otherwise. -/ @@ -60,6 +59,7 @@ section Select variable (P : ℕ → Prop) +open scoped Classical in /-- The polynomial that witnesses that `WittVector.select` is a polynomial function. `selectPoly n` is `X n` if `P n` holds, and `0` otherwise. -/ def selectPoly (n : ℕ) : MvPolynomial ℕ ℤ := diff --git a/Mathlib/Topology/EMetricSpace/Diam.lean b/Mathlib/Topology/EMetricSpace/Diam.lean index 7b5eda92a2ba8..f632f6d65eb0e 100644 --- a/Mathlib/Topology/EMetricSpace/Diam.lean +++ b/Mathlib/Topology/EMetricSpace/Diam.lean @@ -11,7 +11,7 @@ import Mathlib.Topology.EMetricSpace.Pi -/ -open Set Filter Classical +open Set Filter open scoped Uniformity Topology Filter NNReal ENNReal Pointwise From 35ac416c02f542b59afea171c9b876334596b05b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Mon, 30 Dec 2024 21:48:32 +0000 Subject: [PATCH 042/189] feat(Topology): generalize universes for (co)limits in `TopCat` (#20347) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In order to construct the (co)limit of a functor ` J ⥤ TopCat`, we no longer assume `J` is a small category. --- Mathlib/Topology/Category/TopCat/Limits/Basic.lean | 10 +++++----- .../Topology/Category/TopCat/Limits/Cofiltered.lean | 2 +- Mathlib/Topology/Category/TopCat/Limits/Products.lean | 2 +- Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean | 2 +- 4 files changed, 8 insertions(+), 8 deletions(-) diff --git a/Mathlib/Topology/Category/TopCat/Limits/Basic.lean b/Mathlib/Topology/Category/TopCat/Limits/Basic.lean index 4ed86b7b0ecfd..98981cae50ee9 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Basic.lean @@ -23,7 +23,7 @@ noncomputable section namespace TopCat -variable {J : Type v} [SmallCategory J] +variable {J : Type v} [Category.{w} J] local notation "forget" => forget TopCat @@ -103,7 +103,7 @@ def limitConeInfiIsLimit (F : J ⥤ TopCat.{max v u}) : IsLimit (limitConeInfi.{ (continuous_iff_coinduced_le.mp (s.π.app j).continuous : _)) · rfl -instance topCat_hasLimitsOfSize : HasLimitsOfSize.{v, v} TopCat.{max v u} where +instance topCat_hasLimitsOfSize : HasLimitsOfSize.{w, v} TopCat.{max v u} where has_limits_of_shape _ := { has_limit := fun F => HasLimit.mk @@ -114,7 +114,7 @@ instance topCat_hasLimits : HasLimits TopCat.{u} := TopCat.topCat_hasLimitsOfSize.{u, u} instance forget_preservesLimitsOfSize : - PreservesLimitsOfSize.{v, v} (forget : TopCat.{max v u} ⥤ _) where + PreservesLimitsOfSize.{w, v} (forget : TopCat.{max v u} ⥤ _) where preservesLimitsOfShape {_} := { preservesLimit := fun {F} => preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v,u} F) @@ -162,7 +162,7 @@ def colimitCoconeIsColimit (F : J ⥤ TopCat.{max v u}) : IsColimit (colimitCoco (continuous_iff_coinduced_le.mp (s.ι.app j).continuous : _)) · rfl -instance topCat_hasColimitsOfSize : HasColimitsOfSize.{v,v} TopCat.{max v u} where +instance topCat_hasColimitsOfSize : HasColimitsOfSize.{w,v} TopCat.{max v u} where has_colimits_of_shape _ := { has_colimit := fun F => HasColimit.mk @@ -173,7 +173,7 @@ instance topCat_hasColimits : HasColimits TopCat.{u} := TopCat.topCat_hasColimitsOfSize.{u, u} instance forget_preservesColimitsOfSize : - PreservesColimitsOfSize.{v, v} (forget : TopCat.{max u v} ⥤ _) where + PreservesColimitsOfSize.{w, v} (forget : TopCat.{max u v} ⥤ _) where preservesColimitsOfShape := { preservesColimit := fun {F} => preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit F) diff --git a/Mathlib/Topology/Category/TopCat/Limits/Cofiltered.lean b/Mathlib/Topology/Category/TopCat/Limits/Cofiltered.lean index 060971c01bcd8..976f5ed4e73db 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Cofiltered.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Cofiltered.lean @@ -29,7 +29,7 @@ namespace TopCat section CofilteredLimit -variable {J : Type v} [SmallCategory J] [IsCofiltered J] (F : J ⥤ TopCat.{max v u}) (C : Cone F) +variable {J : Type v} [Category.{w} J] [IsCofiltered J] (F : J ⥤ TopCat.{max v u}) (C : Cone F) /-- 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 diff --git a/Mathlib/Topology/Category/TopCat/Limits/Products.lean b/Mathlib/Topology/Category/TopCat/Limits/Products.lean index cb99ab63740be..6512a54157e40 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Products.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Products.lean @@ -22,7 +22,7 @@ noncomputable section namespace TopCat -variable {J : Type v} [SmallCategory J] +variable {J : Type v} [Category.{w} J] /-- The projection from the product as a bundled continuous map. -/ abbrev piπ {ι : Type v} (α : ι → TopCat.{max v u}) (i : ι) : TopCat.of (∀ i, α i) ⟶ α i := diff --git a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean index cbeef70a14d69..f7dfb76f0da23 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean @@ -21,7 +21,7 @@ noncomputable section namespace TopCat -variable {J : Type v} [SmallCategory J] +variable {J : Type v} [Category.{w} J] section Pullback From 5bf0f3993d81f682aab73c6129b25b49da7144fa Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Mon, 30 Dec 2024 22:02:03 +0000 Subject: [PATCH 043/189] chore: rename `Pi.fintype` to `Pi.instFintype` (#7850) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/Pi.20instance.20names.20in.20Lean.204.20core) Co-authored-by: Yaël Dillies Co-authored-by: Eric Wieser --- Mathlib/Combinatorics/SimpleGraph/Finsubgraph.lean | 2 +- Mathlib/Data/Fintype/Pi.lean | 2 +- Mathlib/FieldTheory/Minpoly/Field.lean | 2 +- Mathlib/InformationTheory/Hamming.lean | 2 +- Mathlib/RingTheory/WittVector/Truncated.lean | 2 +- 5 files changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/Combinatorics/SimpleGraph/Finsubgraph.lean b/Mathlib/Combinatorics/SimpleGraph/Finsubgraph.lean index bb66154688f83..d536e4b7b0dcc 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Finsubgraph.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Finsubgraph.lean @@ -157,7 +157,7 @@ theorem nonempty_hom_of_forall_finite_subgraph_hom [Finite W] haveI : ∀ G' : G.Finsubgraphᵒᵖ, Fintype ((finsubgraphHomFunctor G F).obj G') := by intro G' haveI : Fintype (G'.unop.val.verts : Type u) := G'.unop.property.fintype - haveI : Fintype (↥G'.unop.val.verts → W) := by classical exact Pi.fintype + haveI : Fintype (↥G'.unop.val.verts → W) := by classical exact Pi.instFintype exact Fintype.ofInjective (fun f => f.toFun) RelHom.coe_fn_injective -- Use compactness to obtain a section. obtain ⟨u, hu⟩ := nonempty_sections_of_finite_inverse_system (finsubgraphHomFunctor G F) diff --git a/Mathlib/Data/Fintype/Pi.lean b/Mathlib/Data/Fintype/Pi.lean index 2414aec167706..2d19a52f707a6 100644 --- a/Mathlib/Data/Fintype/Pi.lean +++ b/Mathlib/Data/Fintype/Pi.lean @@ -136,7 +136,7 @@ end Fintype /-! ### pi -/ /-- A dependent product of fintypes, indexed by a fintype, is a fintype. -/ -instance Pi.fintype {α : Type*} {β : α → Type*} [DecidableEq α] [Fintype α] +instance Pi.instFintype {α : Type*} {β : α → Type*} [DecidableEq α] [Fintype α] [∀ a, Fintype (β a)] : Fintype (∀ a, β a) := ⟨Fintype.piFinset fun _ => univ, by simp⟩ diff --git a/Mathlib/FieldTheory/Minpoly/Field.lean b/Mathlib/FieldTheory/Minpoly/Field.lean index e071240005dbf..0652b50888c6b 100644 --- a/Mathlib/FieldTheory/Minpoly/Field.lean +++ b/Mathlib/FieldTheory/Minpoly/Field.lean @@ -184,7 +184,7 @@ section AlgHomFintype /-- A technical finiteness result. -/ noncomputable def Fintype.subtypeProd {E : Type*} {X : Set E} (hX : X.Finite) {L : Type*} (F : E → Multiset L) : Fintype (∀ x : X, { l : L // l ∈ F x }) := - @Pi.fintype _ _ _ (Finite.fintype hX) _ + @Pi.instFintype _ _ _ (Finite.fintype hX) _ variable (F E K : Type*) [Field F] [Ring E] [CommRing K] [IsDomain K] [Algebra F E] [Algebra F K] [FiniteDimensional F E] diff --git a/Mathlib/InformationTheory/Hamming.lean b/Mathlib/InformationTheory/Hamming.lean index ce87bc6fec535..8ad9b91993a9b 100644 --- a/Mathlib/InformationTheory/Hamming.lean +++ b/Mathlib/InformationTheory/Hamming.lean @@ -215,7 +215,7 @@ instance [∀ i, Inhabited (β i)] : Inhabited (Hamming β) := ⟨fun _ => default⟩ instance [DecidableEq ι] [Fintype ι] [∀ i, Fintype (β i)] : Fintype (Hamming β) := - Pi.fintype + Pi.instFintype instance [Inhabited ι] [∀ i, Nonempty (β i)] [Nontrivial (β default)] : Nontrivial (Hamming β) := Pi.nontrivial diff --git a/Mathlib/RingTheory/WittVector/Truncated.lean b/Mathlib/RingTheory/WittVector/Truncated.lean index 4fa06dc3e127b..cb39757374dc6 100644 --- a/Mathlib/RingTheory/WittVector/Truncated.lean +++ b/Mathlib/RingTheory/WittVector/Truncated.lean @@ -382,7 +382,7 @@ end section Fintype instance {R : Type*} [Fintype R] : Fintype (TruncatedWittVector p n R) := - Pi.fintype + Pi.instFintype variable (p n R) From 8d758ea28e1713f171d3e53d6b5e159d7aac1736 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Tue, 31 Dec 2024 11:21:32 +0000 Subject: [PATCH 044/189] feat(Topology): add `Dense.induction` (#20342) This can be used as ```lean induction x using hs.induction with | mem x hx => sorry | isClosed => sorry ``` when `hs : Dense s` Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib/Topology/Basic.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index 9547c9e03c05f..76073103460ca 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -570,6 +570,12 @@ theorem dense_compl_singleton_iff_not_open : obtain rfl : U = {x} := eq_singleton_iff_nonempty_unique_mem.2 ⟨hne, hUx⟩ exact ho hU +/-- If a closed property holds for a dense subset, it holds for the whole space. -/ +@[elab_as_elim] +lemma Dense.induction (hs : Dense s) {P : X → Prop} + (mem : ∀ x ∈ s, P x) (isClosed : IsClosed { x | P x }) (x : X) : P x := + hs.closure_eq.symm.subset.trans (isClosed.closure_subset_iff.mpr mem) trivial + theorem IsOpen.subset_interior_closure {s : Set X} (s_open : IsOpen s) : s ⊆ interior (closure s) := s_open.subset_interior_iff.mpr subset_closure From 62f34f74926691347828380db8fc67f2676a5110 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Tue, 31 Dec 2024 11:45:52 +0000 Subject: [PATCH 045/189] feat(RingTheory): lemmas about annihilator (#20341) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib/RingTheory/Ideal/Maps.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index 95aa7008f3ea2..a0cca39024554 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -722,6 +722,21 @@ theorem LinearEquiv.annihilator_eq (e : M ≃ₗ[R] M') : Module.annihilator R M = Module.annihilator R M' := (e.annihilator_le_of_surjective e.surjective).antisymm (e.annihilator_le_of_injective e.injective) +theorem Module.comap_annihilator {R₀} [CommSemiring R₀] [Module R₀ M] + [Algebra R₀ R] [IsScalarTower R₀ R M] : + (Module.annihilator R M).comap (algebraMap R₀ R) = Module.annihilator R₀ M := by + ext x + simp [mem_annihilator] + +lemma Module.annihilator_eq_bot {R M} [Ring R] [AddCommGroup M] [Module R M] : + Module.annihilator R M = ⊥ ↔ FaithfulSMul R M := by + rw [← le_bot_iff] + refine ⟨fun H ↦ ⟨fun {r s} H' ↦ ?_⟩, fun ⟨H⟩ {a} ha ↦ ?_⟩ + · rw [← sub_eq_zero] + exact H (Module.mem_annihilator (r := r - s).mpr + (by simp only [sub_smul, H', sub_self, implies_true])) + · exact @H a 0 (by simp [Module.mem_annihilator.mp ha]) + namespace Submodule /-- `N.annihilator` is the ideal of all elements `r : R` such that `r • N = 0`. -/ From e92b32904d55c48d5bda87edfa398e0f401bc018 Mon Sep 17 00:00:00 2001 From: Li Xuanji Date: Tue, 31 Dec 2024 12:20:47 +0000 Subject: [PATCH 046/189] feat: Prove `r_one_zpow` (#20338) Prove `DihedralGroup.r_one_zpow`, which extends `r_one_pow` to the integers. --- Mathlib/GroupTheory/SpecificGroups/Dihedral.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean b/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean index 3618172616b02..443bfc56811c7 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean @@ -90,6 +90,14 @@ theorem sr_mul_r (i j : ZMod n) : sr i * r j = sr (i + j) := theorem sr_mul_sr (i j : ZMod n) : sr i * sr j = r (j - i) := rfl +@[simp] +theorem inv_r (i : ZMod n) : (r i)⁻¹ = r (-i) := + rfl + +@[simp] +theorem inv_sr (i : ZMod n) : (sr i)⁻¹ = sr i := + rfl + theorem one_def : (1 : DihedralGroup n) = r 0 := rfl @@ -134,6 +142,10 @@ theorem r_one_pow (k : ℕ) : (r 1 : DihedralGroup n) ^ k = r k := by norm_cast rw [Nat.one_add] +@[simp] +theorem r_one_zpow (k : ℤ) : (r 1 : DihedralGroup n) ^ k = r k := by + cases k <;> simp + -- @[simp] -- Porting note: simp changes the goal to `r 0 = 1`. `r_one_pow_n` is no longer useful. theorem r_one_pow_n : r (1 : ZMod n) ^ n = 1 := by rw [r_one_pow, one_def] From 0e87cb2cf55c760c8e5f7621ef33a6aa5a2e32d2 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 31 Dec 2024 16:08:18 +0000 Subject: [PATCH 047/189] feat(Normed/Group): add `tendsto_norm_div_self_nhdsGE` (#20311) --- Mathlib/Analysis/Normed/Group/Basic.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index e69d345223589..ba1878a99efcb 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -849,6 +849,10 @@ theorem tendsto_norm_div_self (x : E) : Tendsto (fun a => ‖a / x‖) (𝓝 x) simpa [dist_eq_norm_div] using tendsto_id.dist (tendsto_const_nhds : Tendsto (fun _a => (x : E)) (𝓝 x) _) +@[to_additive] +theorem tendsto_norm_div_self_nhdsGE (x : E) : Tendsto (fun a ↦ ‖a / x‖) (𝓝 x) (𝓝[≥] 0) := + tendsto_nhdsWithin_iff.mpr ⟨tendsto_norm_div_self x, by simp⟩ + @[to_additive tendsto_norm] theorem tendsto_norm' {x : E} : Tendsto (fun a => ‖a‖) (𝓝 x) (𝓝 ‖x‖) := by simpa using tendsto_id.dist (tendsto_const_nhds : Tendsto (fun _a => (1 : E)) _ _) From bd452dbb5347ff7f5fc68ed1e1a3fe03cbfd83ea Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 31 Dec 2024 16:28:30 +0000 Subject: [PATCH 048/189] chore: relocate `RefinedDiscrTree` (#20308) This should shrink the diff of another open PR --- Mathlib.lean | 2 +- Mathlib/{Tactic/FunProp => Lean/Meta}/RefinedDiscrTree.lean | 0 Mathlib/Tactic.lean | 1 - Mathlib/Tactic/FunProp.lean | 2 +- Mathlib/Tactic/FunProp/Theorems.lean | 2 +- scripts/noshake.json | 2 +- 6 files changed, 4 insertions(+), 5 deletions(-) rename Mathlib/{Tactic/FunProp => Lean/Meta}/RefinedDiscrTree.lean (100%) diff --git a/Mathlib.lean b/Mathlib.lean index 9b2281b6fb61a..030e0edb7cb1a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3320,6 +3320,7 @@ import Mathlib.Lean.Meta.Basic import Mathlib.Lean.Meta.CongrTheorems import Mathlib.Lean.Meta.DiscrTree import Mathlib.Lean.Meta.KAbstractPositions +import Mathlib.Lean.Meta.RefinedDiscrTree import Mathlib.Lean.Meta.Simp import Mathlib.Lean.Name import Mathlib.Lean.PrettyPrinter.Delaborator @@ -4845,7 +4846,6 @@ import Mathlib.Tactic.FunProp.Differentiable import Mathlib.Tactic.FunProp.Elab import Mathlib.Tactic.FunProp.FunctionData import Mathlib.Tactic.FunProp.Mor -import Mathlib.Tactic.FunProp.RefinedDiscrTree import Mathlib.Tactic.FunProp.StateList import Mathlib.Tactic.FunProp.Theorems import Mathlib.Tactic.FunProp.ToBatteries diff --git a/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean b/Mathlib/Lean/Meta/RefinedDiscrTree.lean similarity index 100% rename from Mathlib/Tactic/FunProp/RefinedDiscrTree.lean rename to Mathlib/Lean/Meta/RefinedDiscrTree.lean diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index 369301daef0a9..301a16c9ca93e 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -94,7 +94,6 @@ import Mathlib.Tactic.FunProp.Differentiable import Mathlib.Tactic.FunProp.Elab import Mathlib.Tactic.FunProp.FunctionData import Mathlib.Tactic.FunProp.Mor -import Mathlib.Tactic.FunProp.RefinedDiscrTree import Mathlib.Tactic.FunProp.StateList import Mathlib.Tactic.FunProp.Theorems import Mathlib.Tactic.FunProp.ToBatteries diff --git a/Mathlib/Tactic/FunProp.lean b/Mathlib/Tactic/FunProp.lean index bc97ea7cc4def..1e9f65989489a 100644 --- a/Mathlib/Tactic/FunProp.lean +++ b/Mathlib/Tactic/FunProp.lean @@ -9,7 +9,7 @@ import Mathlib.Tactic.FunProp.Decl import Mathlib.Tactic.FunProp.Elab import Mathlib.Tactic.FunProp.FunctionData import Mathlib.Tactic.FunProp.Mor -import Mathlib.Tactic.FunProp.RefinedDiscrTree +import Mathlib.Lean.Meta.RefinedDiscrTree import Mathlib.Tactic.FunProp.StateList import Mathlib.Tactic.FunProp.Theorems import Mathlib.Tactic.FunProp.ToBatteries diff --git a/Mathlib/Tactic/FunProp/Theorems.lean b/Mathlib/Tactic/FunProp/Theorems.lean index 0ed63b6296c34..abbe43c45aac5 100644 --- a/Mathlib/Tactic/FunProp/Theorems.lean +++ b/Mathlib/Tactic/FunProp/Theorems.lean @@ -6,7 +6,7 @@ Authors: Tomáš Skřivan import Mathlib.Tactic.FunProp.Decl import Mathlib.Tactic.FunProp.Types import Mathlib.Tactic.FunProp.FunctionData -import Mathlib.Tactic.FunProp.RefinedDiscrTree +import Mathlib.Lean.Meta.RefinedDiscrTree import Batteries.Data.RBMap.Alter /-! diff --git a/scripts/noshake.json b/scripts/noshake.json index 6339d0b82139c..22f216add9304 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -253,7 +253,6 @@ "Mathlib.Tactic.GCongr.Core": ["Mathlib.Order.Defs", "Mathlib.Order.Defs.PartialOrder"], "Mathlib.Tactic.GCongr": ["Mathlib.Tactic.Positivity.Core"], - "Mathlib.Tactic.FunProp.RefinedDiscrTree": ["Mathlib.Algebra.Group.Pi.Basic"], "Mathlib.Tactic.FunProp.Mor": ["Mathlib.Data.FunLike.Basic"], "Mathlib.Tactic.FunProp.Differentiable": ["Mathlib.Analysis.Calculus.Deriv.Inv", @@ -359,6 +358,7 @@ "Mathlib.LinearAlgebra.AffineSpace.FiniteDimensional": ["Mathlib.Init.Core"], "Mathlib.LinearAlgebra.AffineSpace.Basic": ["Mathlib.Algebra.AddTorsor"], "Mathlib.Lean.Meta": ["Batteries.Logic"], + "Mathlib.Lean.Meta.RefinedDiscrTree": ["Mathlib.Algebra.Group.Pi.Basic"], "Mathlib.Lean.Expr.ExtraRecognizers": ["Mathlib.Data.Set.Operations"], "Mathlib.Lean.Expr.Basic": ["Batteries.Logic"], "Mathlib.GroupTheory.MonoidLocalization.Basic": ["Mathlib.Init.Data.Prod"], From 584f39cda63a699d6f7e417321885d5e452aa304 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Tue, 31 Dec 2024 19:27:58 +0000 Subject: [PATCH 049/189] feat: `Countable` deriving handler (#18557) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Example: ```lean inductive PList (α : Sort _) where | nil | cons (x : α) (xs : PList α) deriving Countable ``` This creates the instance ```lean instance {α : Sort*} [Countable α] : Countable (PList α) ``` The deriving handler supports structures and all non-reflexive un-nested inductive types, including mutual inductive types. --- Mathlib.lean | 1 + Mathlib/Tactic.lean | 1 + Mathlib/Tactic/DeriveCountable.lean | 276 ++++++++++++++++++++++++++++ MathlibTest/deriving_countable.lean | 102 ++++++++++ 4 files changed, 380 insertions(+) create mode 100644 Mathlib/Tactic/DeriveCountable.lean create mode 100644 MathlibTest/deriving_countable.lean diff --git a/Mathlib.lean b/Mathlib.lean index 030e0edb7cb1a..0e9233cd3add4 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4818,6 +4818,7 @@ import Mathlib.Tactic.Core import Mathlib.Tactic.DeclarationNames import Mathlib.Tactic.DefEqTransformations import Mathlib.Tactic.DeprecateTo +import Mathlib.Tactic.DeriveCountable import Mathlib.Tactic.DeriveFintype import Mathlib.Tactic.DeriveToExpr import Mathlib.Tactic.DeriveTraversable diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index 301a16c9ca93e..b9bfe92da053b 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -66,6 +66,7 @@ import Mathlib.Tactic.Core import Mathlib.Tactic.DeclarationNames import Mathlib.Tactic.DefEqTransformations import Mathlib.Tactic.DeprecateTo +import Mathlib.Tactic.DeriveCountable import Mathlib.Tactic.DeriveFintype import Mathlib.Tactic.DeriveToExpr import Mathlib.Tactic.DeriveTraversable diff --git a/Mathlib/Tactic/DeriveCountable.lean b/Mathlib/Tactic/DeriveCountable.lean new file mode 100644 index 0000000000000..9b15d1388f40a --- /dev/null +++ b/Mathlib/Tactic/DeriveCountable.lean @@ -0,0 +1,276 @@ +/- +Copyright (c) 2024 Kyle Miller. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kyle Miller +-/ +import Lean.Meta.Transform +import Lean.Meta.Inductive +import Lean.Elab.Deriving.Basic +import Lean.Elab.Deriving.Util +import Mathlib.Data.Countable.Defs +import Mathlib.Data.Nat.Pairing + +/-! +# `Countable` deriving handler + +Adds a deriving handler for the `Countable` class. +-/ + +namespace Mathlib.Deriving.Countable +open Lean Parser.Term Elab Deriving Meta + +/-! +### Theory + +We use the `Nat.pair` function to encode an inductive type in the natural numbers, +following a pattern similar to the tagged s-expressions used in Scheme/Lisp. +We develop a little theory to make constructing the injectivity functions very straightforward. + +This is easiest to explain by example. Given a type +```lean +inductive T (α : Type) + | a (n : α) + | b (n m : α) (t : T α) +``` +the deriving handler constructs the following three declarations: +```lean +noncomputable def T.toNat (α : Type) [Countable α] : T α → ℕ + | a n => Nat.pair 0 (Nat.pair (encode n) 0) + | b n m t => Nat.pair 1 (Nat.pair (encode n) (Nat.pair (encode m) (Nat.pair t.toNat 0))) + +theorem T.toNat_injective (α : Type) [Countable α] : Function.Injective (T.toNat α) := fun + | a .., a .. => by + refine cons_eq_imp_init ?_ + refine pair_encode_step ?_; rintro ⟨⟩ + intro; rfl + | a .., b .. => by refine cons_eq_imp ?_; rintro ⟨⟩ + | b .., a .. => by refine cons_eq_imp ?_; rintro ⟨⟩ + | b .., b .. => by + refine cons_eq_imp_init ?_ + refine pair_encode_step ?_; rintro ⟨⟩ + refine pair_encode_step ?_; rintro ⟨⟩ + refine cons_eq_imp ?_; intro h; cases T.toNat_injective α h + intro; rfl + +instance (α : Type) [Countable α] : Countable (T α) := ⟨_, T.toNat_injective α⟩ +``` +-/ + +private noncomputable def encode {α : Sort*} [Countable α] : α → ℕ := + (Countable.exists_injective_nat α).choose + +private noncomputable def encode_injective {α : Sort*} [Countable α] : + Function.Injective (encode : α → ℕ) := + (Countable.exists_injective_nat α).choose_spec + +/-- +Initialize the injectivity argument. Pops the constructor tag. +-/ +private theorem cons_eq_imp_init {p : Prop} {a b b' : ℕ} + (h : b = b' → p) : Nat.pair a b = Nat.pair a b' → p := by + simpa [Nat.pair_eq_pair, and_imp] using h + +/-- +Generic step of the injectivity argument, pops the head of the pairs. +Used in the recursive steps where we need to supply an additional injectivity argument. +-/ +private theorem cons_eq_imp {p : Prop} {a b a' b' : ℕ} + (h : a = a' → b = b' → p) : Nat.pair a b = Nat.pair a' b' → p := by + rwa [Nat.pair_eq_pair, and_imp] + +/-- +Specialized step of the injectivity argument, pops the head of the pairs and decodes. +-/ +private theorem pair_encode_step {p : Prop} {α : Sort*} [Countable α] + {a b : α} {m n : ℕ} + (h : a = b → m = n → p) : Nat.pair (encode a) m = Nat.pair (encode b) n → p := + cons_eq_imp fun ha => h (encode_injective ha) + + +/-! +### Implementation +-/ + +/-! +Constructing the `toNat` functions. +-/ + +private def mkToNatMatch (ctx : Deriving.Context) (header : Header) (indVal : InductiveVal) + (toFunNames : Array Name) : TermElabM Term := do + let discrs ← mkDiscrs header indVal + let alts ← mkAlts + `(match $[$discrs],* with $alts:matchAlt*) +where + mkAlts : TermElabM (Array (TSyntax ``matchAlt)) := do + let mut alts := #[] + for ctorName in indVal.ctors do + let ctorInfo ← getConstInfoCtor ctorName + alts := alts.push <| ← forallTelescopeReducing ctorInfo.type fun xs _ => do + let mut patterns := #[] + let mut ctorArgs := #[] + let mut rhsArgs : Array Term := #[] + for _ in [:indVal.numIndices] do + patterns := patterns.push (← `(_)) + for _ in [:ctorInfo.numParams] do + ctorArgs := ctorArgs.push (← `(_)) + for i in [:ctorInfo.numFields] do + let a := mkIdent (← mkFreshUserName `a) + ctorArgs := ctorArgs.push a + let x := xs[ctorInfo.numParams + i]! + let xTy ← inferType x + let recName? := ctx.typeInfos.findIdx? (xTy.isAppOf ·.name) |>.map (toFunNames[·]!) + rhsArgs := rhsArgs.push <| ← + if let some recName := recName? then + `($(mkIdent recName) $a) + else + ``(encode $a) + patterns := patterns.push (← `(@$(mkIdent ctorName):ident $ctorArgs:term*)) + let rhs' : Term ← rhsArgs.foldrM (init := ← `(0)) fun arg acc => ``(Nat.pair $arg $acc) + let rhs : Term ← ``(Nat.pair $(quote ctorInfo.cidx) $rhs') + `(matchAltExpr| | $[$patterns:term],* => $rhs) + return alts + +/-- Constructs a function from the inductive type to `Nat`. -/ +def mkToNatFuns (ctx : Deriving.Context) (toNatFnNames : Array Name) : + TermElabM (TSyntax `command) := do + let mut res : Array (TSyntax `command) := #[] + for i in [:toNatFnNames.size] do + let toNatFnName := toNatFnNames[i]! + let indVal := ctx.typeInfos[i]! + let header ← mkHeader ``Countable 1 indVal + let body ← mkToNatMatch ctx header indVal toNatFnNames + res := res.push <| ← `( + private noncomputable def $(mkIdent toNatFnName):ident $header.binders:bracketedBinder* : + Nat := $body:term + ) + `(command| mutual $[$res:command]* end) + +/-! +Constructing the injectivity proofs for these `toNat` functions. +-/ + +private def mkInjThmMatch (ctx : Deriving.Context) (header : Header) (indVal : InductiveVal) : + TermElabM Term := do + let discrs ← mkDiscrs header indVal + let alts ← mkAlts + `(match $[$discrs],* with $alts:matchAlt*) +where + mkAlts : TermElabM (Array (TSyntax ``matchAlt)) := do + let mut alts := #[] + for ctorName₁ in indVal.ctors do + let ctorInfo ← getConstInfoCtor ctorName₁ + for ctorName₂ in indVal.ctors do + let mut patterns := #[] + for _ in [:indVal.numIndices] do + patterns := patterns.push (← `(_)) + patterns := patterns ++ #[(← `($(mkIdent ctorName₁) ..)), (← `($(mkIdent ctorName₂) ..))] + if ctorName₁ == ctorName₂ then + alts := alts.push <| ← forallTelescopeReducing ctorInfo.type fun xs _ => do + let mut tactics : Array (TSyntax `tactic) := #[] + for i in [:ctorInfo.numFields] do + let x := xs[indVal.numParams + i]! + let xTy ← inferType x + let recName? := + ctx.typeInfos.findIdx? (xTy.isAppOf ·.name) |>.map (ctx.auxFunNames[·]!) + tactics := tactics.push <| ← + if let some recName := recName? then + `(tactic| ( + refine $(mkCIdent ``cons_eq_imp) ?_; + intro h; + cases $(mkIdent recName) _ _ h + )) + else + `(tactic| ( + refine $(mkCIdent ``pair_encode_step) ?_; + rintro ⟨⟩ + )) + tactics := tactics.push (← `(tactic| (intro; rfl))) + `(matchAltExpr| | $[$patterns:term],* => cons_eq_imp_init (by $tactics:tactic*)) + else if (← compatibleCtors ctorName₁ ctorName₂) then + let rhs ← ``(cons_eq_imp (by rintro ⟨⟩)) + alts := alts.push (← `(matchAltExpr| | $[$patterns:term],* => $rhs:term)) + return alts + +/-- Constructs a proof that the functions created by `mkToNatFuns` are injective. -/ +def mkInjThms (ctx : Deriving.Context) (toNatFnNames : Array Name) : + TermElabM (TSyntax `command) := do + let mut res : Array (TSyntax `command) := #[] + for i in [:toNatFnNames.size] do + let toNatFnName := toNatFnNames[i]! + let injThmName := ctx.auxFunNames[i]! + let indVal := ctx.typeInfos[i]! + let header ← mkHeader ``Countable 2 indVal + let body ← mkInjThmMatch ctx header indVal + let f := mkIdent toNatFnName + let t1 := mkIdent header.targetNames[0]! + let t2 := mkIdent header.targetNames[1]! + res := res.push <| ← `( + private theorem $(mkIdent injThmName):ident $header.binders:bracketedBinder* : + $f $t1 = $f $t2 → $t1 = $t2 := $body:term + ) + `(command| mutual $[$res:command]* end) + +/-! +Assembling the `Countable` instances. +-/ + +open TSyntax.Compat in +/-- Assuming all of the auxiliary definitions exist, create all the `instance` commands +for the `ToExpr` instances for the (mutual) inductive type(s). -/ +private def mkCountableInstanceCmds (ctx : Deriving.Context) (typeNames : Array Name) : + TermElabM (Array Command) := do + let mut instances := #[] + for i in [:ctx.typeInfos.size] do + let indVal := ctx.typeInfos[i]! + if typeNames.contains indVal.name then + let auxFunName := ctx.auxFunNames[i]! + let argNames ← mkInductArgNames indVal + let binders ← mkImplicitBinders argNames + let binders := binders ++ (← mkInstImplicitBinders ``Countable indVal argNames) + let indType ← mkInductiveApp indVal argNames + let type ← `($(mkCIdent ``Countable) $indType) + let mut val := mkIdent auxFunName + let instCmd ← `(instance $binders:implicitBinder* : $type := ⟨_, $val⟩) + instances := instances.push instCmd + return instances + +private def mkCountableCmds (indVal : InductiveVal) (declNames : Array Name) : + TermElabM (Array Syntax) := do + let ctx ← mkContext "countable" indVal.name + let toNatFunNames : Array Name ← ctx.auxFunNames.mapM fun name => do + let .str n' s := name.eraseMacroScopes | unreachable! + mkFreshUserName <| .str n' (s ++ "ToNat") + let cmds := #[← mkToNatFuns ctx toNatFunNames, ← mkInjThms ctx toNatFunNames] + ++ (← mkCountableInstanceCmds ctx declNames) + trace[Mathlib.Deriving.countable] "\n{cmds}" + return cmds + +open Command + +/-- +The deriving handler for the `Countable` class. +Handles non-nested non-reflexive inductive types. +They can be mutual too — in that case, there is an optimization to re-use all the generated +functions and proofs. +-/ +def mkCountableInstance (declNames : Array Name) : CommandElabM Bool := do + let mut seen : NameSet := {} + let mut toVisit : Array InductiveVal := #[] + for declName in declNames do + if seen.contains declName then continue + let indVal ← getConstInfoInduct declName + if indVal.isNested || indVal.isReflexive then + return false -- not supported yet + seen := seen.append (NameSet.ofList indVal.all) + toVisit := toVisit.push indVal + for indVal in toVisit do + let cmds ← liftTermElabM <| mkCountableCmds indVal (declNames.filter indVal.all.contains) + withEnableInfoTree false do + elabCommand <| mkNullNode cmds + return true + +initialize + registerDerivingHandler ``Countable mkCountableInstance + registerTraceClass `Mathlib.Deriving.countable + +end Mathlib.Deriving.Countable diff --git a/MathlibTest/deriving_countable.lean b/MathlibTest/deriving_countable.lean new file mode 100644 index 0000000000000..e1e9b233f4cdd --- /dev/null +++ b/MathlibTest/deriving_countable.lean @@ -0,0 +1,102 @@ +import Mathlib.Tactic.DeriveCountable +import Mathlib.Data.Countable.Basic +import Mathlib.Logic.Equiv.List + +/-! +# Tests of the `Countable` deriving handler +-/ + +namespace Test + +/-! +Structure +-/ +structure S (α β : Type*) where + x : α + y : List β + deriving Countable + +example : Countable (S Nat Bool) := inferInstance + +/-! +Recursive inductive type, no indices +-/ +inductive T (α : Type) + | a (n : Option α) + | b (n m : α) (t : T α) + | c + deriving Countable + +example : Countable (T Nat) := inferInstance + +/-! +Motivating example for writing the deriving handler +-/ +inductive PreCantor' : Type + | zero : PreCantor' + | oadd : PreCantor' → ℕ+ → PreCantor' → PreCantor' + deriving Countable + +example : Countable PreCantor' := inferInstance + +/-! +Sort +-/ +inductive PList (α : Sort _) where + | nil + | cons (x : α) (xs : PList α) + deriving Countable + +example : Countable (PList (1 = 2)) := inferInstance + +/-! +Recursive type with indices +-/ + +inductive I : Nat → Type + | a (n : Nat) : I n + | b (n : Nat) (x : Nat) (c : I x) : I (n + 1) + deriving Countable + +example : Countable (I 37) := inferInstance +example : Countable (Σ i, I i) := inferInstance + +/-! +Mutually recursive types +-/ +mutual +inductive T1 where + | a + | b (x : T1) (y : T2) + deriving Countable +inductive T2 where + | c (n : Nat) + | d (x y : T1) + deriving Countable +end + +example : Countable T1 := inferInstance +example : Countable T2 := inferInstance + + +/-! +Not supported: nested inductive types +-/ +/-- +error: default handlers have not been implemented yet, class: 'Countable' types: [Test.Nested] +-/ +#guard_msgs in +inductive Nested where + | mk (xs : List Nested) + deriving Countable + +/-! +Not supported: reflexive inductive types +-/ +/-- +error: default handlers have not been implemented yet, class: 'Countable' types: [Test.Reflex] +-/ +#guard_msgs in +inductive Reflex where + | mk (f : Bool → Reflex) + deriving Countable From 7178aee7a431bb7527da15c3507836d8dfefcda4 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Tue, 31 Dec 2024 21:01:59 +0000 Subject: [PATCH 050/189] chore(RingTheory/Ideal/Colon): generalize to semiring (#18595) Keep the annihilator eq as a lemma for CommRing --- Mathlib/RingTheory/Ideal/Colon.lean | 52 +++++++++++++++++++++++------ 1 file changed, 42 insertions(+), 10 deletions(-) diff --git a/Mathlib/RingTheory/Ideal/Colon.lean b/Mathlib/RingTheory/Ideal/Colon.lean index 20098bc349739..37b90aaaf0c8f 100644 --- a/Mathlib/RingTheory/Ideal/Colon.lean +++ b/Mathlib/RingTheory/Ideal/Colon.lean @@ -18,17 +18,32 @@ namespace Submodule open Pointwise -variable {R M M' F G : Type*} [CommRing R] [AddCommGroup M] [Module R M] -variable {N N₁ N₂ P P₁ P₂ : Submodule R M} +variable {R M M' F G : Type*} + +section Semiring + +variable [Semiring R] [AddCommMonoid M] [Module R M] +variable {N P : Submodule R M} /-- `N.colon P` is the ideal of all elements `r : R` such that `r • P ⊆ N`. -/ -def colon (N P : Submodule R M) : Ideal R := - annihilator (P.map N.mkQ) +def colon (N P : Submodule R M) : Ideal R where + carrier := {r : R | (r • P : Set M) ⊆ N} + add_mem' ha hb := + (Set.add_smul_subset _ _ _).trans ((Set.add_subset_add ha hb).trans_eq (by simp [← coe_sup])) + zero_mem' := by simp [Set.zero_smul_set P.nonempty] + smul_mem' r := by + simp only [Set.mem_setOf_eq, smul_eq_mul, mul_smul, Set.smul_set_subset_iff] + intro x hx y hy + exact N.smul_mem _ (hx hy) + +theorem mem_colon {r} : r ∈ N.colon P ↔ ∀ p ∈ P, r • p ∈ N := Set.smul_set_subset_iff -theorem mem_colon {r} : r ∈ N.colon P ↔ ∀ p ∈ P, r • p ∈ N := - mem_annihilator.trans - ⟨fun H p hp => (Quotient.mk_eq_zero N).1 (H (Quotient.mk p) (mem_map_of_mem hp)), - fun H _ ⟨p, hp, hpm⟩ => hpm ▸ ((Quotient.mk_eq_zero N).2 <| H p hp)⟩ +end Semiring + +section CommSemiring + +variable [CommSemiring R] [AddCommMonoid M] [Module R M] +variable {N N₁ N₂ P P₁ P₂ : Submodule R M} theorem mem_colon' {r} : r ∈ N.colon P ↔ P ≤ comap (r • (LinearMap.id : M →ₗ[R] M)) N := mem_colon @@ -71,12 +86,29 @@ theorem _root_.Ideal.mem_colon_singleton {I : Ideal R} {x r : R} : r ∈ I.colon (Ideal.span {x}) ↔ r * x ∈ I := by simp only [← Ideal.submodule_span_eq, Submodule.mem_colon_singleton, smul_eq_mul] +end CommSemiring + +section CommRing + +variable [CommRing R] [AddCommGroup M] [Module R M] +variable {N N₁ N₂ P P₁ P₂ : Submodule R M} + +@[simp] +lemma annihilator_map_mkQ_eq_colon : annihilator (P.map N.mkQ) = N.colon P := by + ext + rw [mem_annihilator, mem_colon] + exact ⟨fun H p hp ↦ (Quotient.mk_eq_zero N).1 (H (Quotient.mk p) (mem_map_of_mem hp)), + fun H _ ⟨p, hp, hpm⟩ ↦ hpm ▸ ((Quotient.mk_eq_zero N).2 <| H p hp)⟩ + theorem annihilator_quotient {N : Submodule R M} : Module.annihilator R (M ⧸ N) = N.colon ⊤ := by - simp [SetLike.ext_iff, Module.mem_annihilator, colon, - LinearMap.range_eq_top.mpr (mkQ_surjective N)] + simp_rw [SetLike.ext_iff, Module.mem_annihilator, ←annihilator_map_mkQ_eq_colon, mem_annihilator, + map_top, LinearMap.range_eq_top.mpr (mkQ_surjective N), mem_top, forall_true_left, + forall_const] theorem _root_.Ideal.annihilator_quotient {I : Ideal R} : Module.annihilator R (R ⧸ I) = I := by rw [Submodule.annihilator_quotient, colon_top] +end CommRing + end Submodule From bd681bc33a22201938feb3796dfad236496726dd Mon Sep 17 00:00:00 2001 From: Li Xuanji Date: Wed, 1 Jan 2025 11:13:53 +0000 Subject: [PATCH 051/189] doc: tidy the docstring of `Subgroup.centralizer` (#20365) The argument is now named `s` as it has type `Set G` (not `Subgroup G`) --- Mathlib/GroupTheory/Subgroup/Centralizer.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Centralizer.lean b/Mathlib/GroupTheory/Subgroup/Centralizer.lean index 9a0fb7f77947e..70ce36b8a26b4 100644 --- a/Mathlib/GroupTheory/Subgroup/Centralizer.lean +++ b/Mathlib/GroupTheory/Subgroup/Centralizer.lean @@ -17,9 +17,9 @@ namespace Subgroup variable {H K : Subgroup G} -/-- The `centralizer` of `H` is the subgroup of `g : G` commuting with every `h : H`. -/ +/-- The `centralizer` of `s` is the subgroup of `g : G` commuting with every `h : s`. -/ @[to_additive - "The `centralizer` of `H` is the additive subgroup of `g : G` commuting with every `h : H`."] + "The `centralizer` of `s` is the additive subgroup of `g : G` commuting with every `h : s`."] def centralizer (s : Set G) : Subgroup G := { Submonoid.centralizer s with carrier := Set.centralizer s From 9419c034110d02f83c8cf286f877c624e492e16f Mon Sep 17 00:00:00 2001 From: smorel394 Date: Wed, 1 Jan 2025 14:50:29 +0000 Subject: [PATCH 052/189] chore(CategoryTheory/Adjunction/Opposites): redefine the opposite of an adjunction (#20375) Change the definition of `Adjunction.adjointOfOpAdjointOp` and `Adjunction.opAdjointOpOfAdjoint` to be more in line with the new definition of adjunctions from #16317 (which removes the `homEquiv` field), and rename these to `Adjunction.op` and `Adjunction.unop` to be consistent with `Equivalence.op` and `Equivalence.unop`. The 6 other variants of opposite adjunctions are all defined from `Adjunction.op` or `Adjunction.unop` and should not be needed anymore, they have been marked as `deprecated`, as well as `Adjunction.adjointOfOpAdjointOp` and `Adjunction.opAdjointOpOfAdjoint`. Co-authored-by: smorel394 <67864981+smorel394@users.noreply.github.com> --- .../CategoryTheory/Adjunction/Opposites.lean | 82 ++++++------------- 1 file changed, 23 insertions(+), 59 deletions(-) diff --git a/Mathlib/CategoryTheory/Adjunction/Opposites.lean b/Mathlib/CategoryTheory/Adjunction/Opposites.lean index 615dd8ae5d899..90cff54fc2a2d 100644 --- a/Mathlib/CategoryTheory/Adjunction/Opposites.lean +++ b/Mathlib/CategoryTheory/Adjunction/Opposites.lean @@ -29,66 +29,31 @@ namespace CategoryTheory.Adjunction attribute [local simp] homEquiv_unit homEquiv_counit -/-- If `G.op` is adjoint to `F.op` then `F` is adjoint to `G`. -/ -@[simps! unit_app counit_app] -def adjointOfOpAdjointOp (F : C ⥤ D) (G : D ⥤ C) (h : G.op ⊣ F.op) : F ⊣ G := - Adjunction.mkOfHomEquiv { - homEquiv := fun {X Y} => - ((h.homEquiv (Opposite.op Y) (Opposite.op X)).trans (opEquiv _ _)).symm.trans - (opEquiv _ _) - homEquiv_naturality_left_symm := by - -- Porting note: This proof was handled by `obviously` in mathlib3. The only obstruction to - -- automation fully kicking in here is that the `@[simps]` lemmas of `opEquiv` and - -- `homEquiv` aren't firing. - intros - simp [opEquiv, homEquiv] - homEquiv_naturality_right := by - -- Porting note: This proof was handled by `obviously` in mathlib3. The only obstruction to - -- automation fully kicking in here is that the `@[simps]` lemmas of `opEquiv` and - -- `homEquiv` aren't firing. - intros - simp [opEquiv, homEquiv] } - -/-- If `G` is adjoint to `F.op` then `F` is adjoint to `G.unop`. -/ -def adjointUnopOfAdjointOp (F : C ⥤ D) (G : Dᵒᵖ ⥤ Cᵒᵖ) (h : G ⊣ F.op) : F ⊣ G.unop := - adjointOfOpAdjointOp F G.unop (h.ofNatIsoLeft G.opUnopIso.symm) - -/-- If `G.op` is adjoint to `F` then `F.unop` is adjoint to `G`. -/ -def unopAdjointOfOpAdjoint (F : Cᵒᵖ ⥤ Dᵒᵖ) (G : D ⥤ C) (h : G.op ⊣ F) : F.unop ⊣ G := - adjointOfOpAdjointOp _ _ (h.ofNatIsoRight F.opUnopIso.symm) - /-- If `G` is adjoint to `F` then `F.unop` is adjoint to `G.unop`. -/ -def unopAdjointUnopOfAdjoint (F : Cᵒᵖ ⥤ Dᵒᵖ) (G : Dᵒᵖ ⥤ Cᵒᵖ) (h : G ⊣ F) : F.unop ⊣ G.unop := - adjointUnopOfAdjointOp F.unop G (h.ofNatIsoRight F.opUnopIso.symm) +@[simps] +def unop {F : Cᵒᵖ ⥤ Dᵒᵖ} {G : Dᵒᵖ ⥤ Cᵒᵖ} (h : G ⊣ F) : F.unop ⊣ G.unop where + unit := NatTrans.unop h.counit + counit := NatTrans.unop h.unit + left_triangle_components _ := Quiver.Hom.op_inj (h.right_triangle_components _) + right_triangle_components _ := Quiver.Hom.op_inj (h.left_triangle_components _) + +@[deprecated (since := "2025-01-01")] alias adjointOfOpAdjointOp := unop +@[deprecated (since := "2025-01-01")] alias adjointUnopOfAdjointOp := unop +@[deprecated (since := "2025-01-01")] alias unopAdjointOfOpAdjoint := unop +@[deprecated (since := "2025-01-01")] alias unopAdjointUnopOfAdjoint := unop /-- If `G` is adjoint to `F` then `F.op` is adjoint to `G.op`. -/ -@[simps! unit_app counit_app] -def opAdjointOpOfAdjoint (F : C ⥤ D) (G : D ⥤ C) (h : G ⊣ F) : F.op ⊣ G.op := - Adjunction.mkOfHomEquiv { - homEquiv := fun X Y => - (opEquiv _ Y).trans ((h.homEquiv _ _).symm.trans (opEquiv X (Opposite.op _)).symm) - homEquiv_naturality_left_symm := by - -- Porting note: This proof was handled by `obviously` in mathlib3. The only obstruction to - -- automation fully kicking in here is that the `@[simps]` lemmas of `opEquiv` aren't firing. - intros - simp [opEquiv] - homEquiv_naturality_right := by - -- Porting note: This proof was handled by `obviously` in mathlib3. The only obstruction to - -- automation fully kicking in here is that the `@[simps]` lemmas of `opEquiv` aren't firing. - intros - simp [opEquiv] } - -/-- If `G` is adjoint to `F.unop` then `F` is adjoint to `G.op`. -/ -def adjointOpOfAdjointUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) (G : D ⥤ C) (h : G ⊣ F.unop) : F ⊣ G.op := - (opAdjointOpOfAdjoint F.unop _ h).ofNatIsoLeft F.opUnopIso - -/-- If `G.unop` is adjoint to `F` then `F.op` is adjoint to `G`. -/ -def opAdjointOfUnopAdjoint (F : C ⥤ D) (G : Dᵒᵖ ⥤ Cᵒᵖ) (h : G.unop ⊣ F) : F.op ⊣ G := - (opAdjointOpOfAdjoint _ G.unop h).ofNatIsoRight G.opUnopIso - -/-- If `G.unop` is adjoint to `F.unop` then `F` is adjoint to `G`. -/ -def adjointOfUnopAdjointUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) (G : Dᵒᵖ ⥤ Cᵒᵖ) (h : G.unop ⊣ F.unop) : F ⊣ G := - (adjointOpOfAdjointUnop _ _ h).ofNatIsoRight G.opUnopIso +@[simps] +def op {F : C ⥤ D} {G : D ⥤ C} (h : G ⊣ F) : F.op ⊣ G.op where + unit := NatTrans.op h.counit + counit := NatTrans.op h.unit + left_triangle_components _ := Quiver.Hom.unop_inj (by simp) + right_triangle_components _ := Quiver.Hom.unop_inj (by simp) + +@[deprecated (since := "2025-01-01")] alias opAdjointOpOfAdjoint := op +@[deprecated (since := "2025-01-01")] alias adjointOpOfAdjointUnop := op +@[deprecated (since := "2025-01-01")] alias opAdjointOfUnopAdjoint := op +@[deprecated (since := "2025-01-01")] alias adjointOfUnopAdjointUnop := op /-- If `F` and `F'` are both adjoint to `G`, there is a natural isomorphism `F.op ⋙ coyoneda ≅ F'.op ⋙ coyoneda`. @@ -119,7 +84,6 @@ Note: it is generally better to use `Adjunction.natIsoEquiv`, see the file `Adju -/ def natIsoOfLeftAdjointNatIso {F F' : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G') (l : F ≅ F') : G ≅ G' := - NatIso.removeOp (natIsoOfRightAdjointNatIso (opAdjointOpOfAdjoint _ F' adj2) - (opAdjointOpOfAdjoint _ _ adj1) (NatIso.op l)) + NatIso.removeOp (natIsoOfRightAdjointNatIso (op adj2) (op adj1) (NatIso.op l)) end CategoryTheory.Adjunction From 504007f6f772de93ea967b17320baeac5a21cb55 Mon Sep 17 00:00:00 2001 From: smorel394 Date: Wed, 1 Jan 2025 17:49:20 +0000 Subject: [PATCH 053/189] feat(CategoryTheory/Shift/Adjunction): commutation with shifts and adjunctions, part 2 (#20273) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given categories `C` and `D` that have shifts by an additive group `A`, functors `F : C ⥤ D` and `G : C ⥤ D`, an adjunction `F ⊣ G` and a `CommShift` structure on `G`, this file constructs a `CommShift` structure on `F`. This is a sequel to PR #20033, which was doing the construction in the other direction. --- Mathlib/CategoryTheory/Shift/Adjunction.lean | 117 ++++++++++++++++--- 1 file changed, 101 insertions(+), 16 deletions(-) diff --git a/Mathlib/CategoryTheory/Shift/Adjunction.lean b/Mathlib/CategoryTheory/Shift/Adjunction.lean index 0cbde40a773ff..ec22e3951bc7b 100644 --- a/Mathlib/CategoryTheory/Shift/Adjunction.lean +++ b/Mathlib/CategoryTheory/Shift/Adjunction.lean @@ -11,11 +11,15 @@ import Mathlib.CategoryTheory.Adjunction.Mates Given categories `C` and `D` that have shifts by an additive group `A`, functors `F : C ⥤ D` and `G : C ⥤ D`, an adjunction `F ⊣ G` and a `CommShift` structure on `F`, this file constructs -a `CommShift` structure on `G`. As an easy application, if `E : C ≌ D` is an equivalence and -`E.functor` has a `CommShift` structure, we get a `CommShift` structure on `E.inverse`. - -The `CommShift` structure on `G` must be compatible with the one on `F` in the following sense -(cf. `Adjunction.CommShift`): +a `CommShift` structure on `G`. We also do the construction in the other direction: given a +`CommShift` structure on `G`, we construct a `CommShift` structure on `G`; we could do this +using opposite categories, but the construction is simple enough that it is not really worth it. +As an easy application, if `E : C ≌ D` is an equivalence and `E.functor` has a `CommShift` +structure, we get a `CommShift` structure on `E.inverse`. + +We now explain the construction of a `CommShift` structure on `G` given a `CommShift` structure +on `F`; the other direction is similar. The `CommShift` structure on `G` must be compatible with +the one on `F` in the following sense (cf. `Adjunction.CommShift`): for every `a` in `A`, the natural transformation `adj.unit : 𝟭 C ⟶ G ⋙ F` commutes with the isomorphism `shiftFunctor C A ⋙ G ⋙ F ≅ G ⋙ F ⋙ shiftFunctor C A` induces by `F.commShiftIso a` and `G.commShiftIso a`. We actually require a similar condition for @@ -37,9 +41,6 @@ Once we have established all this, the compatibility of the commutation isomorph `F` expressed in `CommShift.zero` and `CommShift.add` immediately implies the similar statements for the commutation isomorphisms for `G`. -TODO: Construct a `CommShift` structure on `F` from a `CommShift` structure on `G`, using -opposite categories. - -/ namespace CategoryTheory @@ -142,12 +143,13 @@ lemma compatibilityUnit_unique_right (h : CompatibilityUnit adj e₁ e₂) /-- Given an adjunction `adj : F ⊣ G`, `a` in `A` and commutation isomorphisms `e₁ : shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a` and `e₂ : shiftFunctor D a ⋙ G ≅ G ⋙ shiftFunctor C a`, if `e₁` and `e₂` are compatible with the -counit of the adjunction `adj`, then `e₂` uniquely determines `e₁`. +unit of the adjunction `adj`, then `e₂` uniquely determines `e₁`. -/ -lemma compatibilityCounit_unique_left (h : CompatibilityCounit adj e₁ e₂) - (h' : CompatibilityCounit adj e₁' e₂) : e₁ = e₁' := by +lemma compatibilityUnit_unique_left (h : CompatibilityUnit adj e₁ e₂) + (h' : CompatibilityUnit adj e₁' e₂) : e₁ = e₁' := by ext - rw [compatibilityCounit_left adj e₁ e₂ h, compatibilityCounit_left adj e₁' e₂ h'] + rw [compatibilityCounit_left adj e₁ e₂ (compatibilityCounit_of_compatibilityUnit adj _ _ h), + compatibilityCounit_left adj e₁' e₂ (compatibilityCounit_of_compatibilityUnit adj _ _ h')] /-- The isomorphisms `Functor.CommShift.isoZero F` and `Functor.CommShift.isoZero G` are @@ -271,8 +273,7 @@ lemma iso_hom_app (X : D) : (G.map ((shiftFunctorCompIsoId D a b (by rw [← add_left_inj a, add_assoc, h, zero_add, add_zero])).hom.app X))⟦a⟧' := by obtain rfl : b = -a := by rw [← add_left_inj a, h, neg_add_cancel] - simp [iso, iso'] - rfl + simp [iso, iso', shiftEquiv'] @[reassoc] lemma iso_inv_app (Y : D) : @@ -327,12 +328,12 @@ the unique compatible `CommShift` structure on `G`. noncomputable def rightAdjointCommShift [F.CommShift A] : G.CommShift A where iso a := iso adj a zero := by - refine CommShift.compatibilityUnit_unique_right adj (F.commShiftIso 0) _ _ + refine CommShift.compatibilityUnit_unique_right adj (F.commShiftIso 0) _ _ (compatibilityUnit_iso adj 0) ?_ rw [F.commShiftIso_zero] exact CommShift.compatibilityUnit_isoZero adj add a b := by - refine CommShift.compatibilityUnit_unique_right adj (F.commShiftIso (a + b)) _ _ + refine CommShift.compatibilityUnit_unique_right adj (F.commShiftIso (a + b)) _ _ (compatibilityUnit_iso adj (a + b)) ?_ rw [F.commShiftIso_add] exact CommShift.compatibilityUnit_isoAdd adj _ _ _ _ @@ -348,6 +349,90 @@ lemma commShift_of_leftAdjoint [F.CommShift A] : simpa only [Functor.commShiftIso_id_hom_app, Functor.comp_obj, Functor.id_obj, id_comp, Functor.commShiftIso_comp_hom_app] using RightAdjointCommShift.compatibilityUnit_iso adj a X +namespace LeftAdjointCommShift + +variable {A} (a b : A) (h : a + b = 0) [G.CommShift A] + +/-- Auxiliary definition for `iso`. -/ +noncomputable def iso' : shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a := + (conjugateIsoEquiv (Adjunction.comp adj (shiftEquiv' D a b h).toAdjunction) + (Adjunction.comp (shiftEquiv' C a b h).toAdjunction adj)).invFun (G.commShiftIso b) + +/-- +Given an adjunction `F ⊣ G` and a `CommShift` structure on `G`, these are the candidate +`CommShift.iso a` isomorphisms for a compatible `CommShift` structure on `F`. +-/ +noncomputable def iso : shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a := + iso' adj _ _ (add_neg_cancel a) + +@[reassoc] +lemma iso_hom_app (X : C) : + (iso adj a).hom.app X = F.map ((adj.unit.app X)⟦a⟧') ≫ + F.map (G.map (((shiftFunctorCompIsoId D a b h).inv.app (F.obj X)))⟦a⟧') ≫ + F.map (((G.commShiftIso b).hom.app ((F.obj X)⟦a⟧))⟦a⟧') ≫ + F.map ((shiftFunctorCompIsoId C b a (by simp [eq_neg_of_add_eq_zero_left h])).hom.app + (G.obj ((F.obj X)⟦a⟧))) ≫ adj.counit.app ((F.obj X)⟦a⟧) := by + obtain rfl : b = -a := eq_neg_of_add_eq_zero_right h + simp [iso, iso', shiftEquiv'] + +@[reassoc] +lemma iso_inv_app (Y : C) : + (iso adj a).inv.app Y = (F.map ((shiftFunctorCompIsoId C a b h).inv.app Y))⟦a⟧' ≫ + (F.map ((adj.unit.app (Y⟦a⟧))⟦b⟧'))⟦a⟧' ≫ (F.map ((G.commShiftIso b).inv.app + (F.obj (Y⟦a⟧))))⟦a⟧' ≫ (adj.counit.app ((F.obj (Y⟦a⟧))⟦b⟧))⟦a⟧' ≫ + (shiftFunctorCompIsoId D b a (by simp [eq_neg_of_add_eq_zero_left h])).hom.app + (F.obj (Y⟦a⟧)) := by + obtain rfl : b = -a := eq_neg_of_add_eq_zero_right h + simp [iso, iso', shiftEquiv'] + +/-- +The commutation isomorphisms of `Adjunction.LeftAdjointCommShift.iso` are compatible with +the unit of the adjunction. +-/ +lemma compatibilityUnit_iso (a : A) : + CommShift.CompatibilityUnit adj (iso adj a) (G.commShiftIso a) := by + intro + rw [LeftAdjointCommShift.iso_hom_app adj _ _ (add_neg_cancel a)] + simp only [Functor.id_obj, Functor.comp_obj, Functor.map_shiftFunctorCompIsoId_inv_app, + Functor.map_comp, assoc, unit_naturality_assoc, right_triangle_components_assoc] + slice_rhs 4 5 => rw [← Functor.map_comp, Iso.inv_hom_id_app] + simp only [Functor.comp_obj, Functor.map_id, id_comp] + rw [shift_shiftFunctorCompIsoId_inv_app, ← Functor.comp_map, + (shiftFunctorCompIsoId C _ _ (neg_add_cancel a)).hom.naturality_assoc] + simp + +end LeftAdjointCommShift + +open LeftAdjointCommShift in +/-- +Given an adjunction `F ⊣ G` and a `CommShift` structure on `G`, this constructs +the unique compatible `CommShift` structure on `F`. +-/ +@[simps] +noncomputable def leftAdjointCommShift [G.CommShift A] : F.CommShift A where + iso a := iso adj a + zero := by + refine CommShift.compatibilityUnit_unique_left adj _ _ (G.commShiftIso 0) + (compatibilityUnit_iso adj 0) ?_ + rw [G.commShiftIso_zero] + exact CommShift.compatibilityUnit_isoZero adj + add a b := by + refine CommShift.compatibilityUnit_unique_left adj _ _ (G.commShiftIso (a + b)) + (compatibilityUnit_iso adj (a + b)) ?_ + rw [G.commShiftIso_add] + exact CommShift.compatibilityUnit_isoAdd adj _ _ _ _ + (compatibilityUnit_iso adj a) (compatibilityUnit_iso adj b) + +lemma commShift_of_rightAdjoint [G.CommShift A] : + letI := adj.leftAdjointCommShift A + adj.CommShift A := by + letI := adj.leftAdjointCommShift A + refine CommShift.mk' _ _ ⟨fun a ↦ ?_⟩ + ext X + dsimp + simpa only [Functor.commShiftIso_id_hom_app, Functor.comp_obj, Functor.id_obj, id_comp, + Functor.commShiftIso_comp_hom_app] using LeftAdjointCommShift.compatibilityUnit_iso adj a X + end Adjunction namespace Equivalence From c781ca864408200f5f68bff153a83bd95b3775da Mon Sep 17 00:00:00 2001 From: smorel394 Date: Wed, 1 Jan 2025 20:42:13 +0000 Subject: [PATCH 054/189] feat(Algebra/Category/Grp/Ulift): some properties of the universe lift functor for groups. (#19968) Prove some basic properties of the universe lift functor for groups: - It is fully faithful. - It preserves all limits and creates small limits. - On the category of additive commutative groups, it preserves zero morphisms and is an additive functor. TODO: colimits. - [x] depends on: #19966 Co-authored-by: smorel394 <67864981+smorel394@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Algebra/Category/Grp/Ulift.lean | 135 ++++++++++++++++++++++++ 2 files changed, 136 insertions(+) create mode 100644 Mathlib/Algebra/Category/Grp/Ulift.lean diff --git a/Mathlib.lean b/Mathlib.lean index 0e9233cd3add4..64737dac22cb7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -88,6 +88,7 @@ import Mathlib.Algebra.Category.Grp.Kernels import Mathlib.Algebra.Category.Grp.Limits import Mathlib.Algebra.Category.Grp.Preadditive import Mathlib.Algebra.Category.Grp.Subobject +import Mathlib.Algebra.Category.Grp.Ulift import Mathlib.Algebra.Category.Grp.ZModuleEquivalence import Mathlib.Algebra.Category.Grp.Zero import Mathlib.Algebra.Category.GrpWithZero diff --git a/Mathlib/Algebra/Category/Grp/Ulift.lean b/Mathlib/Algebra/Category/Grp/Ulift.lean new file mode 100644 index 0000000000000..4e58356a103e1 --- /dev/null +++ b/Mathlib/Algebra/Category/Grp/Ulift.lean @@ -0,0 +1,135 @@ +/- +Copyright (c) 2024 Sophie Morel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sophie Morel +-/ +import Mathlib.Algebra.Category.Grp.Limits +import Mathlib.CategoryTheory.Limits.Preserves.Ulift +import Mathlib.CategoryTheory.Preadditive.AdditiveFunctor + +/-! +# Properties of the universe lift functor for groups + +This file shows that the functors `Grp.uliftFunctor` and `CommGrp.uliftFunctor` +(as well as the additive versions) are fully faithful, preserves all limits and +create small limits. + +It also shows that `AddCommGrp.uliftFunctor` is an additive functor. + +-/ + +universe v w w' u + +open CategoryTheory Limits + +namespace Grp + +/-- The universe lift functor for groups is fully faithful. +-/ +@[to_additive + "The universe lift functor for additive groups is fully faithful."] +def uliftFunctorFullyFaithful : uliftFunctor.{u, v}.FullyFaithful where + preimage f := Grp.ofHom (MulEquiv.ulift.toMonoidHom.comp + (f.comp MulEquiv.ulift.symm.toMonoidHom)) + map_preimage _ := rfl + preimage_map _ := rfl + +/-- The universe lift functor for groups is faithful. +-/ +@[to_additive + "The universe lift functor for additive groups is faithful."] +instance : uliftFunctor.{u, v}.Faithful := uliftFunctorFullyFaithful.faithful + + +/-- The universe lift functor for groups is full. +-/ +@[to_additive + "The universe lift functor for additive groups is full."] +instance : uliftFunctor.{u, v}.Full := uliftFunctorFullyFaithful.full + +@[to_additive] +noncomputable instance uliftFunctor_preservesLimit {J : Type w} [Category.{w'} J] + (K : J ⥤ Grp.{u}) : PreservesLimit K uliftFunctor.{v, u} where + preserves lc := ⟨isLimitOfReflects (forget Grp.{max u v}) + (isLimitOfPreserves CategoryTheory.uliftFunctor (isLimitOfPreserves (forget Grp) lc))⟩ + +@[to_additive] +noncomputable instance uliftFunctor_preservesLimitsOfShape {J : Type w} [Category.{w'} J] : + PreservesLimitsOfShape J uliftFunctor.{v, u} where + +/-- +The universe lift for groups preserves limits of arbitrary size. +-/ +@[to_additive + "The universe lift functor for additive groups preserves limits of arbitrary size."] +noncomputable instance uliftFunctor_preservesLimitsOfSize : + PreservesLimitsOfSize.{w', w} uliftFunctor.{v, u} where + +/-- +The universe lift functor on `Grp.{u}` creates `u`-small limits. +-/ +@[to_additive + "The universe lift functor on `AddGrp.{u}` creates `u`-small limits."] +noncomputable instance : CreatesLimitsOfSize.{w, u} uliftFunctor.{v, u} where + CreatesLimitsOfShape := { CreatesLimit := fun {_} ↦ createsLimitOfFullyFaithfulOfPreserves } + +end Grp + +namespace CommGrp + +/-- The universe lift functor for commutative groups is fully faithful. +-/ +@[to_additive + "The universe lift functor for commutative additive groups is fully faithful."] +def uliftFunctorFullyFaithful : uliftFunctor.{u, v}.FullyFaithful where + preimage f := Grp.ofHom (MulEquiv.ulift.toMonoidHom.comp + (f.comp MulEquiv.ulift.symm.toMonoidHom)) + map_preimage _ := rfl + preimage_map _ := rfl + +-- The universe lift functor for commutative groups is faithful. -/ +@[to_additive + "The universe lift functor for commutative additive groups is faithful."] +instance : uliftFunctor.{u, v}.Faithful := uliftFunctorFullyFaithful.faithful + +-- The universe lift functor for commutative groups is full. -/ +@[to_additive + "The universe lift functor for commutative additive groups is full."] +instance : uliftFunctor.{u, v}.Full := uliftFunctorFullyFaithful.full + +@[to_additive] +noncomputable instance uliftFunctor_preservesLimit {J : Type w} [Category.{w'} J] + (K : J ⥤ CommGrp.{u}) : PreservesLimit K uliftFunctor.{v, u} where + preserves lc := ⟨isLimitOfReflects (forget CommGrp.{max u v}) + (isLimitOfPreserves CategoryTheory.uliftFunctor (isLimitOfPreserves (forget CommGrp) lc))⟩ + +@[to_additive] +noncomputable instance uliftFunctor_preservesLimitsOfShape {J : Type w} [Category.{w'} J] : + PreservesLimitsOfShape J uliftFunctor.{v, u} where + +/-- +The universe lift for commutative groups preserves limits of arbitrary size. +-/ +@[to_additive + "The universe lift functor for commutative additive groups preserves limits of arbitrary size."] +noncomputable instance uliftFunctor_preservesLimitsOfSize : + PreservesLimitsOfSize.{w', w} uliftFunctor.{v, u} where + +/-- +The universe lift functor on `CommGrp.{u}` creates `u`-small limits. +-/ +@[to_additive + "The universe lift functor on `AddCommGrp.{u}` creates `u`-small limits."] +noncomputable instance : CreatesLimitsOfSize.{w, u} uliftFunctor.{v, u} where + CreatesLimitsOfShape := { CreatesLimit := fun {_} ↦ createsLimitOfFullyFaithfulOfPreserves } + +end CommGrp + +namespace AddCommGroup + +/-- The universe lift for commutative additive groups is additive. +-/ +instance uliftFunctor_additive : + AddCommGrp.uliftFunctor.{u,v}.Additive where + +end AddCommGroup From 9232fbe19e65abd4066b2a06f038f9082d19eec1 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 1 Jan 2025 21:42:29 +0000 Subject: [PATCH 055/189] refactor: add an ofNat() elaborator (#20169) `ofNat(n)` is a shorthand for `no_index (OfNat.ofNat n)`. Its purpose is: * To make it easier to remember to write the `no_index` * To add a place to put a hoverable docstrings explaining the complexities * To make it easier to remove the `no_index`es in future if `DiscrTree`s stops needing them around `ofNat` (https://github.com/leanprover/lean4/issues/2867). This is not exhaustive, but is a step in the right direction. For now, I only target (some of) the declarations with a `See note [no_index around OfNat.ofNat]`. Some theorem statements have been corrected in passing. This exposed two Lean bugs relating to not using `Expr.consumeMData`: * leanprover/lean4#6438 * leanprover/lean4#6467 These are made more likely by this PR adding `no_index` to the RHS of equalities, but were already possible to trigger by using `simp [<- _]` on existing theorems with `no_index` on the LHS. --- Counterexamples/Cyclotomic105.lean | 5 +- Mathlib.lean | 1 + Mathlib/Algebra/CharP/Defs.lean | 4 +- Mathlib/Algebra/DirectSum/Ring.lean | 3 +- Mathlib/Algebra/Group/Defs.lean | 4 +- Mathlib/Algebra/Group/Hom/End.lean | 3 +- Mathlib/Algebra/Group/Opposite.lean | 6 +- Mathlib/Algebra/Group/ULift.lean | 6 +- Mathlib/Algebra/Module/NatInt.lean | 3 +- Mathlib/Algebra/MvPolynomial/Basic.lean | 15 ++--- Mathlib/Algebra/Order/Floor.lean | 63 +++++++------------ .../Order/Monoid/Unbundled/WithTop.lean | 34 +++++----- Mathlib/Algebra/Polynomial/Basic.lean | 6 +- Mathlib/Algebra/Polynomial/Coeff.lean | 6 +- .../Polynomial/Degree/Definitions.lean | 3 +- Mathlib/Algebra/Ring/Center.lean | 3 +- Mathlib/Algebra/Ring/Int/Parity.lean | 3 +- Mathlib/Data/Fin/Basic.lean | 3 +- Mathlib/Data/Int/Cast/Basic.lean | 3 +- Mathlib/Data/Int/Log.lean | 6 +- Mathlib/Data/Int/Sqrt.lean | 3 +- Mathlib/Data/Nat/Cast/Defs.lean | 6 +- Mathlib/Data/Nat/Cast/Order/Ring.lean | 12 ++-- Mathlib/Data/Nat/Cast/Prod.lean | 6 +- Mathlib/Data/Nat/Cast/Synonym.lean | 6 +- Mathlib/Data/Rat/Cast/Defs.lean | 6 +- Mathlib/Data/Rat/Defs.lean | 11 ++-- Mathlib/Lean/Expr/Basic.lean | 1 + .../Transcendental/Liouville/Residual.lean | 7 ++- Mathlib/Tactic.lean | 1 + Mathlib/Tactic/Common.lean | 1 + Mathlib/Tactic/OfNat.lean | 18 ++++++ Mathlib/Tactic/Rify.lean | 3 +- .../Topology/ContinuousMap/Bounded/Basic.lean | 3 +- scripts/noshake.json | 1 + 35 files changed, 119 insertions(+), 146 deletions(-) create mode 100644 Mathlib/Tactic/OfNat.lean diff --git a/Counterexamples/Cyclotomic105.lean b/Counterexamples/Cyclotomic105.lean index 7d9d4d28b945b..0b0b6c3f5a89a 100644 --- a/Counterexamples/Cyclotomic105.lean +++ b/Counterexamples/Cyclotomic105.lean @@ -94,7 +94,10 @@ theorem cyclotomic_105 : repeat' norm_num theorem coeff_cyclotomic_105 : coeff (cyclotomic 105 ℤ) 7 = -2 := by - simp [cyclotomic_105, coeff_one, coeff_X_of_ne_one] + rw [cyclotomic_105] + -- `Int.reduceNeg` produces an invalid proof that incorrectly transforms this goal into `False`, + -- c.f. https://github.com/leanprover/lean4/issues/6467. + simp [coeff_X_of_ne_one, coeff_one, -Int.reduceNeg] theorem not_forall_coeff_cyclotomic_neg_one_zero_one : ¬∀ n i, coeff (cyclotomic n ℤ) i ∈ ({-1, 0, 1} : Set ℤ) := by diff --git a/Mathlib.lean b/Mathlib.lean index 64737dac22cb7..2de875ff969de 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4941,6 +4941,7 @@ import Mathlib.Tactic.NormNum.Prime import Mathlib.Tactic.NormNum.Result import Mathlib.Tactic.NthRewrite import Mathlib.Tactic.Observe +import Mathlib.Tactic.OfNat import Mathlib.Tactic.PPWithUniv import Mathlib.Tactic.Peel import Mathlib.Tactic.Polyrith diff --git a/Mathlib/Algebra/CharP/Defs.lean b/Mathlib/Algebra/CharP/Defs.lean index 00eef1858e846..0e34907443410 100644 --- a/Mathlib/Algebra/CharP/Defs.lean +++ b/Mathlib/Algebra/CharP/Defs.lean @@ -61,14 +61,12 @@ lemma natCast_eq_natCast' (h : a ≡ b [MOD p]) : (a : R) = b := by @[simp] lemma cast_eq_zero : (p : R) = 0 := (cast_eq_zero_iff R p p).2 dvd_rfl --- See note [no_index around OfNat.ofNat] --- -- TODO: This lemma needs to be `@[simp]` for confluence in the presence of `CharP.cast_eq_zero` and -- `Nat.cast_ofNat`, but with `no_index` on its entire LHS, it matches literally every expression so -- is too expensive. If https://github.com/leanprover/lean4/issues/2867 is fixed in a performant way, this can be made `@[simp]`. -- -- @[simp] -lemma ofNat_eq_zero [p.AtLeastTwo] : no_index (OfNat.ofNat p : R) = 0 := cast_eq_zero R p +lemma ofNat_eq_zero [p.AtLeastTwo] : (ofNat(p) : R) = 0 := cast_eq_zero R p lemma natCast_eq_natCast_mod (a : ℕ) : (a : R) = a % p := natCast_eq_natCast' R p (Nat.mod_modEq a p).symm diff --git a/Mathlib/Algebra/DirectSum/Ring.lean b/Mathlib/Algebra/DirectSum/Ring.lean index 883dc0bb086a7..1ad1ff8412c36 100644 --- a/Mathlib/Algebra/DirectSum/Ring.lean +++ b/Mathlib/Algebra/DirectSum/Ring.lean @@ -424,9 +424,8 @@ instance : NatCast (A 0) := theorem of_natCast (n : ℕ) : of A 0 n = n := rfl --- See note [no_index around OfNat.ofNat] @[simp] -theorem of_zero_ofNat (n : ℕ) [n.AtLeastTwo] : of A 0 (no_index (OfNat.ofNat n)) = OfNat.ofNat n := +theorem of_zero_ofNat (n : ℕ) [n.AtLeastTwo] : of A 0 ofNat(n) = ofNat(n) := of_natCast A n /-- The `Semiring` structure derived from `GSemiring A`. -/ diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index ab65c89102d88..324b405f14be4 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -9,6 +9,7 @@ import Mathlib.Algebra.Group.ZeroOne import Mathlib.Algebra.Group.Operations import Mathlib.Logic.Function.Defs import Mathlib.Tactic.Simps.Basic +import Mathlib.Tactic.OfNat import Batteries.Logic /-! @@ -888,9 +889,8 @@ theorem zpow_natCast (a : G) : ∀ n : ℕ, a ^ (n : ℤ) = a ^ n @[deprecated (since := "2024-03-20")] alias zpow_coe_nat := zpow_natCast @[deprecated (since := "2024-03-20")] alias coe_nat_zsmul := natCast_zsmul --- See note [no_index around OfNat.ofNat] @[to_additive ofNat_zsmul] -lemma zpow_ofNat (a : G) (n : ℕ) : a ^ (no_index (OfNat.ofNat n) : ℤ) = a ^ OfNat.ofNat n := +lemma zpow_ofNat (a : G) (n : ℕ) : a ^ (ofNat(n) : ℤ) = a ^ OfNat.ofNat n := zpow_natCast .. theorem zpow_negSucc (a : G) (n : ℕ) : a ^ (Int.negSucc n) = (a ^ (n + 1))⁻¹ := by diff --git a/Mathlib/Algebra/Group/Hom/End.lean b/Mathlib/Algebra/Group/Hom/End.lean index caa7de5333f3d..f3538c92cd2d9 100644 --- a/Mathlib/Algebra/Group/Hom/End.lean +++ b/Mathlib/Algebra/Group/Hom/End.lean @@ -34,9 +34,8 @@ instance instAddMonoidWithOne (M) [AddCommMonoid M] : AddMonoidWithOne (AddMonoi @[simp] lemma natCast_apply [AddCommMonoid M] (n : ℕ) (m : M) : (↑n : AddMonoid.End M) m = n • m := rfl --- See note [no_index around OfNat.ofNat] @[simp] lemma ofNat_apply [AddCommMonoid M] (n : ℕ) [n.AtLeastTwo] (m : M) : - (no_index (OfNat.ofNat n : AddMonoid.End M)) m = n • m := rfl + (ofNat(n) : AddMonoid.End M) m = n • m := rfl instance instSemiring [AddCommMonoid M] : Semiring (AddMonoid.End M) := { AddMonoid.End.monoid M, AddMonoidHom.addCommMonoid, AddMonoid.End.instAddMonoidWithOne M with diff --git a/Mathlib/Algebra/Group/Opposite.lean b/Mathlib/Algebra/Group/Opposite.lean index fbb777e3f9af9..527a665d95d21 100644 --- a/Mathlib/Algebra/Group/Opposite.lean +++ b/Mathlib/Algebra/Group/Opposite.lean @@ -208,10 +208,9 @@ end DivInvMonoid theorem op_natCast [NatCast α] (n : ℕ) : op (n : α) = n := rfl --- See note [no_index around OfNat.ofNat] @[to_additive (attr := simp)] theorem op_ofNat [NatCast α] (n : ℕ) [n.AtLeastTwo] : - op (no_index (OfNat.ofNat n : α)) = OfNat.ofNat n := + op (ofNat(n) : α) = ofNat(n) := rfl @[to_additive (attr := simp, norm_cast)] @@ -222,10 +221,9 @@ theorem op_intCast [IntCast α] (n : ℤ) : op (n : α) = n := theorem unop_natCast [NatCast α] (n : ℕ) : unop (n : αᵐᵒᵖ) = n := rfl --- See note [no_index around OfNat.ofNat] @[to_additive (attr := simp)] theorem unop_ofNat [NatCast α] (n : ℕ) [n.AtLeastTwo] : - unop (no_index (OfNat.ofNat n : αᵐᵒᵖ)) = OfNat.ofNat n := + unop (ofNat(n) : αᵐᵒᵖ) = ofNat(n) := rfl @[to_additive (attr := simp, norm_cast)] diff --git a/Mathlib/Algebra/Group/ULift.lean b/Mathlib/Algebra/Group/ULift.lean index ce331c7da8288..f7088335116e8 100644 --- a/Mathlib/Algebra/Group/ULift.lean +++ b/Mathlib/Algebra/Group/ULift.lean @@ -112,10 +112,9 @@ instance instIntCast [IntCast α] : IntCast (ULift α) := ⟨(up ·)⟩ theorem up_natCast [NatCast α] (n : ℕ) : up (n : α) = n := rfl --- See note [no_index around OfNat.ofNat] @[simp] theorem up_ofNat [NatCast α] (n : ℕ) [n.AtLeastTwo] : - up (no_index (OfNat.ofNat n : α)) = OfNat.ofNat n := + up (ofNat(n) : α) = ofNat(n) := rfl @[simp, norm_cast] @@ -126,10 +125,9 @@ theorem up_intCast [IntCast α] (n : ℤ) : up (n : α) = n := theorem down_natCast [NatCast α] (n : ℕ) : down (n : ULift α) = n := rfl --- See note [no_index around OfNat.ofNat] @[simp] theorem down_ofNat [NatCast α] (n : ℕ) [n.AtLeastTwo] : - down (no_index (OfNat.ofNat n : ULift α)) = OfNat.ofNat n := + down (ofNat(n) : ULift α) = ofNat(n) := rfl @[simp, norm_cast] diff --git a/Mathlib/Algebra/Module/NatInt.lean b/Mathlib/Algebra/Module/NatInt.lean index 22860f1391265..a0ef8033c90e5 100644 --- a/Mathlib/Algebra/Module/NatInt.lean +++ b/Mathlib/Algebra/Module/NatInt.lean @@ -102,9 +102,8 @@ lemma Nat.cast_smul_eq_nsmul (n : ℕ) (b : M) : (n : R) • b = n • b := by | succ n ih => rw [Nat.cast_succ, add_smul, add_smul, one_smul, ih, one_smul] /-- `nsmul` is equal to any other module structure via a cast. -/ --- See note [no_index around OfNat.ofNat] lemma ofNat_smul_eq_nsmul (n : ℕ) [n.AtLeastTwo] (b : M) : - (no_index (OfNat.ofNat n) : R) • b = OfNat.ofNat n • b := Nat.cast_smul_eq_nsmul .. + (ofNat(n) : R) • b = OfNat.ofNat n • b := Nat.cast_smul_eq_nsmul .. /-- `nsmul` is equal to any other module structure via a cast. -/ @[deprecated Nat.cast_smul_eq_nsmul (since := "2024-07-23")] diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index ffbf49a56e5dd..154517230d906 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -922,9 +922,8 @@ theorem eval₂_one : (1 : MvPolynomial σ R).eval₂ f g = 1 := @[simp] theorem eval₂_natCast (n : Nat) : (n : MvPolynomial σ R).eval₂ f g = n := (eval₂_C _ _ _).trans (map_natCast f n) --- See note [no_index around OfNat.ofNat] @[simp] theorem eval₂_ofNat (n : Nat) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n) : MvPolynomial σ R).eval₂ f g = OfNat.ofNat n := + (ofNat(n) : MvPolynomial σ R).eval₂ f g = ofNat(n) := eval₂_natCast f g n @[simp] @@ -1077,9 +1076,8 @@ theorem eval_C : ∀ a, eval f (C a) = a := theorem eval_X : ∀ n, eval f (X n) = f n := eval₂_X _ _ --- See note [no_index around OfNat.ofNat] @[simp] theorem eval_ofNat (n : Nat) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n) : MvPolynomial σ R).eval f = OfNat.ofNat n := + (ofNat(n) : MvPolynomial σ R).eval f = ofNat(n) := map_ofNat _ n @[simp] @@ -1142,9 +1140,8 @@ theorem map_monomial (s : σ →₀ ℕ) (a : R) : map f (monomial s a) = monomi theorem map_C : ∀ a : R, map f (C a : MvPolynomial σ R) = C (f a) := map_monomial _ _ --- See note [no_index around OfNat.ofNat] @[simp] protected theorem map_ofNat (n : Nat) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n) : MvPolynomial σ R).map f = OfNat.ofNat n := + (ofNat(n) : MvPolynomial σ R).map f = ofNat(n) := _root_.map_ofNat _ _ @[simp] @@ -1366,9 +1363,8 @@ theorem aeval_X (s : σ) : aeval f (X s : MvPolynomial _ R) = f s := theorem aeval_C (r : R) : aeval f (C r) = algebraMap R S₁ r := eval₂_C _ _ _ --- See note [no_index around OfNat.ofNat] @[simp] theorem aeval_ofNat (n : Nat) [n.AtLeastTwo] : - aeval f (no_index (OfNat.ofNat n) : MvPolynomial σ R) = OfNat.ofNat n := + aeval f (ofNat(n) : MvPolynomial σ R) = ofNat(n) := map_ofNat _ _ theorem aeval_unique (φ : MvPolynomial σ R →ₐ[R] S₁) : φ = aeval (φ ∘ X) := by @@ -1502,10 +1498,9 @@ theorem aevalTower_X (i : σ) : aevalTower g y (X i) = y i := theorem aevalTower_C (x : R) : aevalTower g y (C x) = g x := eval₂_C _ _ _ --- See note [no_index around OfNat.ofNat] @[simp] theorem aevalTower_ofNat (n : Nat) [n.AtLeastTwo] : - aevalTower g y (no_index (OfNat.ofNat n) : MvPolynomial σ R) = OfNat.ofNat n := + aevalTower g y (ofNat(n) : MvPolynomial σ R) = ofNat(n) := _root_.map_ofNat _ _ @[simp] diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index bc330a5f70b73..a4bf99e945b6e 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -170,9 +170,8 @@ theorem floor_zero : ⌊(0 : α)⌋₊ = 0 := by rw [← Nat.cast_zero, floor_na @[simp] theorem floor_one : ⌊(1 : α)⌋₊ = 1 := by rw [← Nat.cast_one, floor_natCast] --- See note [no_index around OfNat.ofNat] @[simp] -theorem floor_ofNat (n : ℕ) [n.AtLeastTwo] : ⌊no_index (OfNat.ofNat n : α)⌋₊ = n := +theorem floor_ofNat (n : ℕ) [n.AtLeastTwo] : ⌊(ofNat(n) : α)⌋₊ = ofNat(n) := Nat.floor_natCast _ theorem floor_of_nonpos (ha : a ≤ 0) : ⌊a⌋₊ = 0 := @@ -291,9 +290,8 @@ theorem ceil_zero : ⌈(0 : α)⌉₊ = 0 := by rw [← Nat.cast_zero, ceil_natC @[simp] theorem ceil_one : ⌈(1 : α)⌉₊ = 1 := by rw [← Nat.cast_one, ceil_natCast] --- See note [no_index around OfNat.ofNat] @[simp] -theorem ceil_ofNat (n : ℕ) [n.AtLeastTwo] : ⌈no_index (OfNat.ofNat n : α)⌉₊ = n := ceil_natCast n +theorem ceil_ofNat (n : ℕ) [n.AtLeastTwo] : ⌈(ofNat(n) : α)⌉₊ = ofNat(n) := ceil_natCast n @[simp] theorem ceil_eq_zero : ⌈a⌉₊ = 0 ↔ a ≤ 0 := by rw [← Nat.le_zero, ceil_le, Nat.cast_zero] @@ -402,9 +400,8 @@ theorem floor_add_one (ha : 0 ≤ a) : ⌊a + 1⌋₊ = ⌊a⌋₊ + 1 := by -- Porting note: broken `convert floor_add_nat ha 1` rw [← cast_one, floor_add_nat ha 1] --- See note [no_index around OfNat.ofNat] theorem floor_add_ofNat (ha : 0 ≤ a) (n : ℕ) [n.AtLeastTwo] : - ⌊a + (no_index (OfNat.ofNat n))⌋₊ = ⌊a⌋₊ + OfNat.ofNat n := + ⌊a + ofNat(n)⌋₊ = ⌊a⌋₊ + ofNat(n) := floor_add_nat ha n @[simp] @@ -422,10 +419,9 @@ theorem floor_sub_nat [Sub α] [OrderedSub α] [ExistsAddOfLE α] (a : α) (n : theorem floor_sub_one [Sub α] [OrderedSub α] [ExistsAddOfLE α] (a : α) : ⌊a - 1⌋₊ = ⌊a⌋₊ - 1 := mod_cast floor_sub_nat a 1 --- See note [no_index around OfNat.ofNat] @[simp] theorem floor_sub_ofNat [Sub α] [OrderedSub α] [ExistsAddOfLE α] (a : α) (n : ℕ) [n.AtLeastTwo] : - ⌊a - (no_index (OfNat.ofNat n))⌋₊ = ⌊a⌋₊ - OfNat.ofNat n := + ⌊a - ofNat(n)⌋₊ = ⌊a⌋₊ - ofNat(n) := floor_sub_nat a n theorem ceil_add_nat (ha : 0 ≤ a) (n : ℕ) : ⌈a + n⌉₊ = ⌈a⌉₊ + n := @@ -441,9 +437,8 @@ theorem ceil_add_one (ha : 0 ≤ a) : ⌈a + 1⌉₊ = ⌈a⌉₊ + 1 := by -- Porting note: broken `convert ceil_add_nat ha 1` rw [cast_one.symm, ceil_add_nat ha 1] --- See note [no_index around OfNat.ofNat] theorem ceil_add_ofNat (ha : 0 ≤ a) (n : ℕ) [n.AtLeastTwo] : - ⌈a + (no_index (OfNat.ofNat n))⌉₊ = ⌈a⌉₊ + OfNat.ofNat n := + ⌈a + ofNat(n)⌉₊ = ⌈a⌉₊ + ofNat(n) := ceil_add_nat ha n @[bound] @@ -488,9 +483,8 @@ theorem floor_div_nat (a : α) (n : ℕ) : ⌊a / n⌋₊ = ⌊a⌋₊ / n := by · exact lt_div_mul_add hn · exact cast_pos.2 hn --- See note [no_index around OfNat.ofNat] theorem floor_div_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : - ⌊a / (no_index (OfNat.ofNat n))⌋₊ = ⌊a⌋₊ / OfNat.ofNat n := + ⌊a / ofNat(n)⌋₊ = ⌊a⌋₊ / ofNat(n) := floor_div_nat a n /-- Natural division is the floor of field division. -/ @@ -699,8 +693,7 @@ theorem floor_zero : ⌊(0 : α)⌋ = 0 := by rw [← cast_zero, floor_intCast] @[simp] theorem floor_one : ⌊(1 : α)⌋ = 1 := by rw [← cast_one, floor_intCast] --- See note [no_index around OfNat.ofNat] -@[simp] theorem floor_ofNat (n : ℕ) [n.AtLeastTwo] : ⌊(no_index (OfNat.ofNat n : α))⌋ = n := +@[simp] theorem floor_ofNat (n : ℕ) [n.AtLeastTwo] : ⌊(ofNat(n) : α)⌋ = ofNat(n) := floor_natCast n @[mono] @@ -743,20 +736,18 @@ theorem floor_int_add (z : ℤ) (a : α) : ⌊↑z + a⌋ = z + ⌊a⌋ := by theorem floor_add_nat (a : α) (n : ℕ) : ⌊a + n⌋ = ⌊a⌋ + n := by rw [← Int.cast_natCast, floor_add_int] --- See note [no_index around OfNat.ofNat] @[simp] theorem floor_add_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : - ⌊a + (no_index (OfNat.ofNat n))⌋ = ⌊a⌋ + OfNat.ofNat n := + ⌊a + ofNat(n)⌋ = ⌊a⌋ + ofNat(n) := floor_add_nat a n @[simp] theorem floor_nat_add (n : ℕ) (a : α) : ⌊↑n + a⌋ = n + ⌊a⌋ := by rw [← Int.cast_natCast, floor_int_add] --- See note [no_index around OfNat.ofNat] @[simp] theorem floor_ofNat_add (n : ℕ) [n.AtLeastTwo] (a : α) : - ⌊(no_index (OfNat.ofNat n)) + a⌋ = OfNat.ofNat n + ⌊a⌋ := + ⌊ofNat(n) + a⌋ = ofNat(n) + ⌊a⌋ := floor_nat_add n a @[simp] @@ -769,10 +760,9 @@ theorem floor_sub_nat (a : α) (n : ℕ) : ⌊a - n⌋ = ⌊a⌋ - n := by @[simp] theorem floor_sub_one (a : α) : ⌊a - 1⌋ = ⌊a⌋ - 1 := mod_cast floor_sub_nat a 1 --- See note [no_index around OfNat.ofNat] @[simp] theorem floor_sub_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : - ⌊a - (no_index (OfNat.ofNat n))⌋ = ⌊a⌋ - OfNat.ofNat n := + ⌊a - ofNat(n)⌋ = ⌊a⌋ - ofNat(n) := floor_sub_nat a n theorem abs_sub_lt_one_of_floor_eq_floor {α : Type*} [LinearOrderedCommRing α] [FloorRing α] @@ -830,10 +820,9 @@ theorem fract_add_nat (a : α) (m : ℕ) : fract (a + m) = fract a := by @[simp] theorem fract_add_one (a : α) : fract (a + 1) = fract a := mod_cast fract_add_nat a 1 --- See note [no_index around OfNat.ofNat] @[simp] theorem fract_add_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : - fract (a + (no_index (OfNat.ofNat n))) = fract a := + fract (a + ofNat(n)) = fract a := fract_add_nat a n @[simp] @@ -845,10 +834,9 @@ theorem fract_nat_add (n : ℕ) (a : α) : fract (↑n + a) = fract a := by rw [ @[simp] theorem fract_one_add (a : α) : fract (1 + a) = fract a := mod_cast fract_nat_add 1 a --- See note [no_index around OfNat.ofNat] @[simp] theorem fract_ofNat_add (n : ℕ) [n.AtLeastTwo] (a : α) : - fract ((no_index (OfNat.ofNat n)) + a) = fract a := + fract (ofNat(n) + a) = fract a := fract_nat_add n a @[simp] @@ -864,10 +852,9 @@ theorem fract_sub_nat (a : α) (n : ℕ) : fract (a - n) = fract a := by @[simp] theorem fract_sub_one (a : α) : fract (a - 1) = fract a := mod_cast fract_sub_nat a 1 --- See note [no_index around OfNat.ofNat] @[simp] theorem fract_sub_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : - fract (a - (no_index (OfNat.ofNat n))) = fract a := + fract (a - ofNat(n)) = fract a := fract_sub_nat a n -- Was a duplicate lemma under a bad name @@ -921,10 +908,9 @@ theorem fract_intCast (z : ℤ) : fract (z : α) = 0 := by @[simp] theorem fract_natCast (n : ℕ) : fract (n : α) = 0 := by simp [fract] --- See note [no_index around OfNat.ofNat] @[simp] theorem fract_ofNat (n : ℕ) [n.AtLeastTwo] : - fract ((no_index (OfNat.ofNat n)) : α) = 0 := + fract (ofNat(n) : α) = 0 := fract_natCast n theorem fract_floor (a : α) : fract (⌊a⌋ : α) = 0 := @@ -1122,9 +1108,8 @@ theorem ceil_intCast (z : ℤ) : ⌈(z : α)⌉ = z := theorem ceil_natCast (n : ℕ) : ⌈(n : α)⌉ = n := eq_of_forall_ge_iff fun a => by rw [ceil_le, ← cast_natCast, cast_le] --- See note [no_index around OfNat.ofNat] @[simp] -theorem ceil_ofNat (n : ℕ) [n.AtLeastTwo] : ⌈(no_index (OfNat.ofNat n : α))⌉ = n := ceil_natCast n +theorem ceil_ofNat (n : ℕ) [n.AtLeastTwo] : ⌈(ofNat(n) : α)⌉ = ofNat(n) := ceil_natCast n theorem ceil_mono : Monotone (ceil : α → ℤ) := gc_ceil_coe.monotone_l @@ -1143,10 +1128,9 @@ theorem ceil_add_one (a : α) : ⌈a + 1⌉ = ⌈a⌉ + 1 := by -- Porting note: broken `convert ceil_add_int a (1 : ℤ)` rw [← ceil_add_int a (1 : ℤ), cast_one] --- See note [no_index around OfNat.ofNat] @[simp] theorem ceil_add_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : - ⌈a + (no_index (OfNat.ofNat n))⌉ = ⌈a⌉ + OfNat.ofNat n := + ⌈a + ofNat(n)⌉ = ⌈a⌉ + ofNat(n) := ceil_add_nat a n @[simp] @@ -1162,10 +1146,9 @@ theorem ceil_sub_nat (a : α) (n : ℕ) : ⌈a - n⌉ = ⌈a⌉ - n := by theorem ceil_sub_one (a : α) : ⌈a - 1⌉ = ⌈a⌉ - 1 := by rw [eq_sub_iff_add_eq, ← ceil_add_one, sub_add_cancel] --- See note [no_index around OfNat.ofNat] @[simp] theorem ceil_sub_ofNat (a : α) (n : ℕ) [n.AtLeastTwo] : - ⌈a - (no_index (OfNat.ofNat n))⌉ = ⌈a⌉ - OfNat.ofNat n := + ⌈a - ofNat(n)⌉ = ⌈a⌉ - ofNat(n) := ceil_sub_nat a n @[bound] @@ -1362,9 +1345,8 @@ theorem round_one : round (1 : α) = 1 := by simp [round] @[simp] theorem round_natCast (n : ℕ) : round (n : α) = n := by simp [round] --- See note [no_index around OfNat.ofNat] @[simp] -theorem round_ofNat (n : ℕ) [n.AtLeastTwo] : round (no_index (OfNat.ofNat n : α)) = n := +theorem round_ofNat (n : ℕ) [n.AtLeastTwo] : round (ofNat(n) : α) = ofNat(n) := round_natCast n @[simp] @@ -1394,20 +1376,18 @@ theorem round_sub_one (a : α) : round (a - 1) = round a - 1 := by theorem round_add_nat (x : α) (y : ℕ) : round (x + y) = round x + y := mod_cast round_add_int x y --- See note [no_index around OfNat.ofNat] @[simp] theorem round_add_ofNat (x : α) (n : ℕ) [n.AtLeastTwo] : - round (x + (no_index (OfNat.ofNat n))) = round x + OfNat.ofNat n := + round (x + ofNat(n)) = round x + ofNat(n) := round_add_nat x n @[simp] theorem round_sub_nat (x : α) (y : ℕ) : round (x - y) = round x - y := mod_cast round_sub_int x y --- See note [no_index around OfNat.ofNat] @[simp] theorem round_sub_ofNat (x : α) (n : ℕ) [n.AtLeastTwo] : - round (x - (no_index (OfNat.ofNat n))) = round x - OfNat.ofNat n := + round (x - ofNat(n)) = round x - ofNat(n) := round_sub_nat x n @[simp] @@ -1418,10 +1398,9 @@ theorem round_int_add (x : α) (y : ℤ) : round ((y : α) + x) = y + round x := theorem round_nat_add (x : α) (y : ℕ) : round ((y : α) + x) = y + round x := by rw [add_comm, round_add_nat, add_comm] --- See note [no_index around OfNat.ofNat] @[simp] theorem round_ofNat_add (n : ℕ) [n.AtLeastTwo] (x : α) : - round ((no_index (OfNat.ofNat n)) + x) = OfNat.ofNat n + round x := + round (ofNat(n) + x) = ofNat(n) + round x := round_nat_add x n theorem abs_sub_round_eq_min (x : α) : |x - round x| = min (fract x) (1 - fract x) := by diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean index 3d47f485ab7e9..b9bc99747bd14 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean @@ -324,31 +324,30 @@ instance addMonoidWithOne : AddMonoidWithOne (WithTop α) := @[deprecated (since := "2024-04-05")] alias nat_ne_top := natCast_ne_top @[deprecated (since := "2024-04-05")] alias top_ne_nat := top_ne_natCast --- See note [no_index around OfNat.ofNat] @[simp] lemma coe_ofNat (n : ℕ) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n : α) : WithTop α) = OfNat.ofNat n := rfl + ((ofNat(n) : α) : WithTop α) = ofNat(n) := rfl @[simp] lemma coe_eq_ofNat (n : ℕ) [n.AtLeastTwo] (m : α) : - (m : WithTop α) = no_index (OfNat.ofNat n) ↔ m = OfNat.ofNat n := + (m : WithTop α) = ofNat(n) ↔ m = ofNat(n) := coe_eq_coe @[simp] lemma ofNat_eq_coe (n : ℕ) [n.AtLeastTwo] (m : α) : - no_index (OfNat.ofNat n) = (m : WithTop α) ↔ OfNat.ofNat n = m := + ofNat(n) = (m : WithTop α) ↔ ofNat(n) = m := coe_eq_coe -@[simp] lemma ofNat_ne_top (n : ℕ) [n.AtLeastTwo] : no_index (OfNat.ofNat n : WithTop α) ≠ ⊤ := +@[simp] lemma ofNat_ne_top (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : WithTop α) ≠ ⊤ := natCast_ne_top n -@[simp] lemma top_ne_ofNat (n : ℕ) [n.AtLeastTwo] : (⊤ : WithTop α) ≠ no_index (OfNat.ofNat n) := +@[simp] lemma top_ne_ofNat (n : ℕ) [n.AtLeastTwo] : (⊤ : WithTop α) ≠ ofNat(n) := top_ne_natCast n @[simp] lemma map_ofNat {f : α → β} (n : ℕ) [n.AtLeastTwo] : - WithTop.map f (no_index (OfNat.ofNat n : WithTop α)) = f (OfNat.ofNat n) := map_coe f n + WithTop.map f (ofNat(n) : WithTop α) = f (ofNat(n)) := map_coe f n @[simp] lemma map_natCast {f : α → β} (n : ℕ) : WithTop.map f (n : WithTop α) = f n := map_coe f n lemma map_eq_ofNat_iff {f : β → α} {n : ℕ} [n.AtLeastTwo] {a : WithTop β} : - a.map f = OfNat.ofNat n ↔ ∃ x, a = .some x ∧ f x = n := map_eq_some_iff + a.map f = ofNat(n) ↔ ∃ x, a = .some x ∧ f x = n := map_eq_some_iff lemma ofNat_eq_map_iff {f : β → α} {n : ℕ} [n.AtLeastTwo] {a : WithTop β} : - OfNat.ofNat n = a.map f ↔ ∃ x, a = .some x ∧ f x = n := some_eq_map_iff + ofNat(n) = a.map f ↔ ∃ x, a = .some x ∧ f x = n := some_eq_map_iff lemma map_eq_natCast_iff {f : β → α} {n : ℕ} {a : WithTop β} : a.map f = n ↔ ∃ x, a = .some x ∧ f x = n := map_eq_some_iff @@ -533,31 +532,30 @@ instance addMonoidWithOne : AddMonoidWithOne (WithBot α) := WithTop.addMonoidWi @[deprecated (since := "2024-04-05")] alias nat_ne_bot := natCast_ne_bot @[deprecated (since := "2024-04-05")] alias bot_ne_nat := bot_ne_natCast --- See note [no_index around OfNat.ofNat] @[simp] lemma coe_ofNat (n : ℕ) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n : α) : WithBot α) = OfNat.ofNat n := rfl + ((ofNat(n) : α) : WithBot α) = ofNat(n) := rfl @[simp] lemma coe_eq_ofNat (n : ℕ) [n.AtLeastTwo] (m : α) : - (m : WithBot α) = no_index (OfNat.ofNat n) ↔ m = OfNat.ofNat n := + (m : WithBot α) = ofNat(n) ↔ m = ofNat(n) := coe_eq_coe @[simp] lemma ofNat_eq_coe (n : ℕ) [n.AtLeastTwo] (m : α) : - no_index (OfNat.ofNat n) = (m : WithBot α) ↔ OfNat.ofNat n = m := + ofNat(n) = (m : WithBot α) ↔ ofNat(n) = m := coe_eq_coe -@[simp] lemma ofNat_ne_bot (n : ℕ) [n.AtLeastTwo] : no_index (OfNat.ofNat n : WithBot α) ≠ ⊥ := +@[simp] lemma ofNat_ne_bot (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : WithBot α) ≠ ⊥ := natCast_ne_bot n -@[simp] lemma bot_ne_ofNat (n : ℕ) [n.AtLeastTwo] : (⊥ : WithBot α) ≠ no_index (OfNat.ofNat n) := +@[simp] lemma bot_ne_ofNat (n : ℕ) [n.AtLeastTwo] : (⊥ : WithBot α) ≠ ofNat(n) := bot_ne_natCast n @[simp] lemma map_ofNat {f : α → β} (n : ℕ) [n.AtLeastTwo] : - WithBot.map f (no_index (OfNat.ofNat n : WithBot α)) = f (OfNat.ofNat n) := map_coe f n + WithBot.map f (ofNat(n) : WithBot α) = f ofNat(n) := map_coe f n @[simp] lemma map_natCast {f : α → β} (n : ℕ) : WithBot.map f (n : WithBot α) = f n := map_coe f n lemma map_eq_ofNat_iff {f : β → α} {n : ℕ} [n.AtLeastTwo] {a : WithBot β} : - a.map f = OfNat.ofNat n ↔ ∃ x, a = .some x ∧ f x = n := map_eq_some_iff + a.map f = ofNat(n) ↔ ∃ x, a = .some x ∧ f x = n := map_eq_some_iff lemma ofNat_eq_map_iff {f : β → α} {n : ℕ} [n.AtLeastTwo] {a : WithBot β} : - OfNat.ofNat n = a.map f ↔ ∃ x, a = .some x ∧ f x = n := some_eq_map_iff + ofNat(n) = a.map f ↔ ∃ x, a = .some x ∧ f x = n := some_eq_map_iff lemma map_eq_natCast_iff {f : β → α} {n : ℕ} {a : WithBot β} : a.map f = n ↔ ∃ x, a = .some x ∧ f x = n := map_eq_some_iff diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index a5757859e7bfa..d88f59b1d6cca 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -667,16 +667,14 @@ theorem coeff_natCast_ite : (Nat.cast m : R[X]).coeff n = ite (n = 0) m 0 := by @[deprecated (since := "2024-04-17")] alias coeff_nat_cast_ite := coeff_natCast_ite --- See note [no_index around OfNat.ofNat] @[simp] theorem coeff_ofNat_zero (a : ℕ) [a.AtLeastTwo] : - coeff (no_index (OfNat.ofNat a : R[X])) 0 = OfNat.ofNat a := + coeff (ofNat(a) : R[X]) 0 = ofNat(a) := coeff_monomial --- See note [no_index around OfNat.ofNat] @[simp] theorem coeff_ofNat_succ (a n : ℕ) [h : a.AtLeastTwo] : - coeff (no_index (OfNat.ofNat a : R[X])) (n + 1) = 0 := by + coeff (ofNat(a) : R[X]) (n + 1) = 0 := by rw [← Nat.cast_ofNat] simp [-Nat.cast_ofNat] diff --git a/Mathlib/Algebra/Polynomial/Coeff.lean b/Mathlib/Algebra/Polynomial/Coeff.lean index 1997d5ce972f0..de59da07ebfc2 100644 --- a/Mathlib/Algebra/Polynomial/Coeff.lean +++ b/Mathlib/Algebra/Polynomial/Coeff.lean @@ -172,13 +172,11 @@ theorem coeff_mul_C (p : R[X]) (n : ℕ) (a : R) : coeff (p * C a) n = coeff p n @[simp] lemma coeff_natCast_mul {a k : ℕ} : coeff ((a : R[X]) * p) k = a * coeff p k := coeff_C_mul _ --- See note [no_index around OfNat.ofNat] @[simp] lemma coeff_mul_ofNat {a k : ℕ} [Nat.AtLeastTwo a] : - coeff (p * (no_index (OfNat.ofNat a) : R[X])) k = coeff p k * OfNat.ofNat a := coeff_mul_C _ _ _ + coeff (p * (ofNat(a) : R[X])) k = coeff p k * ofNat(a) := coeff_mul_C _ _ _ --- See note [no_index around OfNat.ofNat] @[simp] lemma coeff_ofNat_mul {a k : ℕ} [Nat.AtLeastTwo a] : - coeff ((no_index (OfNat.ofNat a) : R[X]) * p) k = OfNat.ofNat a * coeff p k := coeff_C_mul _ + coeff ((ofNat(a) : R[X]) * p) k = ofNat(a) * coeff p k := coeff_C_mul _ @[simp] lemma coeff_mul_intCast [Ring S] {p : S[X]} {a : ℤ} {k : ℕ} : coeff (p * (a : S[X])) k = coeff p k * (↑a : S) := coeff_mul_C _ _ _ diff --git a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean index 3dabffa2ed70d..a7f16ee7d9aa6 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean @@ -177,10 +177,9 @@ theorem natDegree_natCast (n : ℕ) : natDegree (n : R[X]) = 0 := by @[deprecated (since := "2024-04-17")] alias natDegree_nat_cast := natDegree_natCast --- See note [no_index around OfNat.ofNat] @[simp] theorem natDegree_ofNat (n : ℕ) [Nat.AtLeastTwo n] : - natDegree (no_index (OfNat.ofNat n : R[X])) = 0 := + natDegree (ofNat(n) : R[X]) = 0 := natDegree_natCast _ theorem degree_natCast_le (n : ℕ) : degree (n : R[X]) ≤ 0 := degree_le_of_natDegree_le (by simp) diff --git a/Mathlib/Algebra/Ring/Center.lean b/Mathlib/Algebra/Ring/Center.lean index 73f4fbcbad95f..828c1c41ecadb 100644 --- a/Mathlib/Algebra/Ring/Center.lean +++ b/Mathlib/Algebra/Ring/Center.lean @@ -37,10 +37,9 @@ theorem natCast_mem_center [NonAssocSemiring M] (n : ℕ) : (n : M) ∈ Set.cent | zero => rw [Nat.cast_zero, mul_zero, mul_zero, mul_zero] | succ n ihn => rw [Nat.cast_succ, mul_add, ihn, mul_add, mul_add, mul_one, mul_one] --- See note [no_index around OfNat.ofNat] @[simp] theorem ofNat_mem_center [NonAssocSemiring M] (n : ℕ) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n)) ∈ Set.center M := + ofNat(n) ∈ Set.center M := natCast_mem_center M n @[simp] diff --git a/Mathlib/Algebra/Ring/Int/Parity.lean b/Mathlib/Algebra/Ring/Int/Parity.lean index 8a6ea4b3ac36d..cd9699824a498 100644 --- a/Mathlib/Algebra/Ring/Int/Parity.lean +++ b/Mathlib/Algebra/Ring/Int/Parity.lean @@ -154,10 +154,9 @@ theorem isSquare_natCast_iff {n : ℕ} : IsSquare (n : ℤ) ↔ IsSquare n := by · exact ⟨x.natAbs, (natAbs_mul_natAbs_eq h.symm).symm⟩ · exact ⟨x, mod_cast h⟩ --- See note [no_index around OfNat.ofNat] @[simp] theorem isSquare_ofNat_iff {n : ℕ} : - IsSquare (no_index (OfNat.ofNat n) : ℤ) ↔ IsSquare (OfNat.ofNat n : ℕ) := + IsSquare (ofNat(n) : ℤ) ↔ IsSquare (ofNat(n) : ℕ) := isSquare_natCast_iff end Int diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 4fcd646907769..11123477cd16a 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -1513,9 +1513,8 @@ theorem coe_natCast_eq_mod (m n : ℕ) [NeZero m] : ((n : Fin m) : ℕ) = n % m := rfl --- See note [no_index around OfNat.ofNat] theorem coe_ofNat_eq_mod (m n : ℕ) [NeZero m] : - ((no_index (OfNat.ofNat n) : Fin m) : ℕ) = OfNat.ofNat n % m := + ((ofNat(n) : Fin m) : ℕ) = ofNat(n) % m := rfl section Mul diff --git a/Mathlib/Data/Int/Cast/Basic.lean b/Mathlib/Data/Int/Cast/Basic.lean index 6885e5e943ea9..ade0aeafd71c2 100644 --- a/Mathlib/Data/Int/Cast/Basic.lean +++ b/Mathlib/Data/Int/Cast/Basic.lean @@ -60,10 +60,9 @@ theorem cast_natCast (n : ℕ) : ((n : ℤ) : R) = n := AddGroupWithOne.intCast_ofNat _ -- expected `n` to be implicit, and `HasLiftT` --- See note [no_index around OfNat.ofNat] @[simp, norm_cast] theorem cast_ofNat (n : ℕ) [n.AtLeastTwo] : - ((no_index (OfNat.ofNat n) : ℤ) : R) = OfNat.ofNat n := by + ((ofNat(n) : ℤ) : R) = ofNat(n) := by simpa only [OfNat.ofNat] using AddGroupWithOne.intCast_ofNat (R := R) n @[simp, norm_cast] diff --git a/Mathlib/Data/Int/Log.lean b/Mathlib/Data/Int/Log.lean index 4d5f3c0e38722..a892892decb01 100644 --- a/Mathlib/Data/Int/Log.lean +++ b/Mathlib/Data/Int/Log.lean @@ -72,10 +72,9 @@ theorem log_natCast (b : ℕ) (n : ℕ) : log b (n : R) = Nat.log b n := by · rw [log_of_one_le_right, Nat.floor_natCast] simp --- See note [no_index around OfNat.ofNat] @[simp] theorem log_ofNat (b : ℕ) (n : ℕ) [n.AtLeastTwo] : - log b (no_index (OfNat.ofNat n : R)) = Nat.log b (OfNat.ofNat n) := + log b (ofNat(n) : R) = Nat.log b ofNat(n) := log_natCast b n theorem log_of_left_le_one {b : ℕ} (hb : b ≤ 1) (r : R) : log b r = 0 := by @@ -221,10 +220,9 @@ theorem clog_natCast (b : ℕ) (n : ℕ) : clog b (n : R) = Nat.clog b n := by · simp [clog_of_right_le_one] · rw [clog_of_one_le_right, (Nat.ceil_eq_iff (Nat.succ_ne_zero n)).mpr] <;> simp --- See note [no_index around OfNat.ofNat] @[simp] theorem clog_ofNat (b : ℕ) (n : ℕ) [n.AtLeastTwo] : - clog b (no_index (OfNat.ofNat n : R)) = Nat.clog b (OfNat.ofNat n) := + clog b (ofNat(n) : R) = Nat.clog b ofNat(n) := clog_natCast b n theorem clog_of_left_le_one {b : ℕ} (hb : b ≤ 1) (r : R) : clog b r = 0 := by diff --git a/Mathlib/Data/Int/Sqrt.lean b/Mathlib/Data/Int/Sqrt.lean index 1c6db061ba8f7..f8cccdb7928fe 100644 --- a/Mathlib/Data/Int/Sqrt.lean +++ b/Mathlib/Data/Int/Sqrt.lean @@ -36,9 +36,8 @@ theorem sqrt_nonneg (n : ℤ) : 0 ≤ sqrt n := @[simp, norm_cast] theorem sqrt_natCast (n : ℕ) : Int.sqrt (n : ℤ) = Nat.sqrt n := by rw [sqrt, toNat_ofNat] --- See note [no_index around OfNat.ofNat] @[simp] -theorem sqrt_ofNat (n : ℕ) : Int.sqrt (no_index (OfNat.ofNat n) : ℤ) = Nat.sqrt (OfNat.ofNat n) := +theorem sqrt_ofNat (n : ℕ) : Int.sqrt ofNat(n) = Nat.sqrt ofNat(n) := sqrt_natCast _ end Int diff --git a/Mathlib/Data/Nat/Cast/Defs.lean b/Mathlib/Data/Nat/Cast/Defs.lean index 5575f08cbd7e7..dce8ade4fbffb 100644 --- a/Mathlib/Data/Nat/Cast/Defs.lean +++ b/Mathlib/Data/Nat/Cast/Defs.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro, Gabriel Ebner -/ import Mathlib.Algebra.Group.Defs import Mathlib.Tactic.SplitIfs +import Mathlib.Tactic.OfNat /-! # Cast of natural numbers @@ -62,11 +63,14 @@ library_note "no_index around OfNat.ofNat" When writing lemmas about `OfNat.ofNat` that assume `Nat.AtLeastTwo`, the term needs to be wrapped in `no_index` so as not to confuse `simp`, as `no_index (OfNat.ofNat n)`. +Rather than referencing this library note, use `ofNat(n)` as a shorthand for +`no_index (OfNat.ofNat n)`. + Some discussion is [on Zulip here](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/.E2.9C.94.20Polynomial.2Ecoeff.20example/near/395438147). -/ @[simp, norm_cast] theorem Nat.cast_ofNat {n : ℕ} [NatCast R] [Nat.AtLeastTwo n] : - (Nat.cast (no_index (OfNat.ofNat n)) : R) = OfNat.ofNat n := rfl + (Nat.cast ofNat(n) : R) = ofNat(n) := rfl @[deprecated Nat.cast_ofNat (since := "2024-12-22")] theorem Nat.cast_eq_ofNat {n : ℕ} [NatCast R] [Nat.AtLeastTwo n] : diff --git a/Mathlib/Data/Nat/Cast/Order/Ring.lean b/Mathlib/Data/Nat/Cast/Order/Ring.lean index 693d3ef4c6797..54269e96efeb0 100644 --- a/Mathlib/Data/Nat/Cast/Order/Ring.lean +++ b/Mathlib/Data/Nat/Cast/Order/Ring.lean @@ -30,10 +30,9 @@ theorem cast_nonneg {α} [OrderedSemiring α] (n : ℕ) : 0 ≤ (n : α) := cast_nonneg' n /-- Specialisation of `Nat.ofNat_nonneg'`, which seems to be easier for Lean to use. -/ --- See note [no_index around OfNat.ofNat] @[simp] theorem ofNat_nonneg {α} [OrderedSemiring α] (n : ℕ) [n.AtLeastTwo] : - 0 ≤ (no_index (OfNat.ofNat n : α)) := + 0 ≤ (ofNat(n) : α) := ofNat_nonneg' n @[simp, norm_cast] @@ -53,16 +52,14 @@ variable [NeZero (1 : α)] theorem cast_pos {α} [OrderedSemiring α] [Nontrivial α] {n : ℕ} : (0 : α) < n ↔ 0 < n := cast_pos' /-- See also `Nat.ofNat_pos`, specialised for an `OrderedSemiring`. -/ --- See note [no_index around OfNat.ofNat] @[simp low] -theorem ofNat_pos' {n : ℕ} [n.AtLeastTwo] : 0 < (no_index (OfNat.ofNat n : α)) := +theorem ofNat_pos' {n : ℕ} [n.AtLeastTwo] : 0 < (ofNat(n) : α) := cast_pos'.mpr (NeZero.pos n) /-- Specialisation of `Nat.ofNat_pos'`, which seems to be easier for Lean to use. -/ --- See note [no_index around OfNat.ofNat] @[simp] theorem ofNat_pos {α} [OrderedSemiring α] [Nontrivial α] {n : ℕ} [n.AtLeastTwo] : - 0 < (no_index (OfNat.ofNat n : α)) := + 0 < (ofNat(n) : α) := ofNat_pos' end Nontrivial @@ -84,10 +81,9 @@ theorem cast_tsub [CanonicallyOrderedCommSemiring α] [Sub α] [OrderedSub α] theorem abs_cast [LinearOrderedRing α] (a : ℕ) : |(a : α)| = a := abs_of_nonneg (cast_nonneg a) --- See note [no_index around OfNat.ofNat] @[simp] theorem abs_ofNat [LinearOrderedRing α] (n : ℕ) [n.AtLeastTwo] : - |(no_index (OfNat.ofNat n : α))| = OfNat.ofNat n := + |(ofNat(n) : α)| = ofNat(n) := abs_cast n lemma mul_le_pow {a : ℕ} (ha : a ≠ 1) (b : ℕ) : diff --git a/Mathlib/Data/Nat/Cast/Prod.lean b/Mathlib/Data/Nat/Cast/Prod.lean index 9903b6c4873ac..1f38821eeb1df 100644 --- a/Mathlib/Data/Nat/Cast/Prod.lean +++ b/Mathlib/Data/Nat/Cast/Prod.lean @@ -26,19 +26,17 @@ instance instAddMonoidWithOne : AddMonoidWithOne (α × β) := @[simp] theorem fst_natCast (n : ℕ) : (n : α × β).fst = n := by induction n <;> simp [*] --- See note [no_index around OfNat.ofNat] @[simp] theorem fst_ofNat (n : ℕ) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n : α × β)).1 = (OfNat.ofNat n : α) := + (ofNat(n) : α × β).1 = (ofNat(n) : α) := rfl @[simp] theorem snd_natCast (n : ℕ) : (n : α × β).snd = n := by induction n <;> simp [*] --- See note [no_index around OfNat.ofNat] @[simp] theorem snd_ofNat (n : ℕ) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n : α × β)).2 = (OfNat.ofNat n : β) := + (ofNat(n) : α × β).2 = (ofNat(n) : β) := rfl end Prod diff --git a/Mathlib/Data/Nat/Cast/Synonym.lean b/Mathlib/Data/Nat/Cast/Synonym.lean index 075d62d0d87cf..11b1f4d3b233a 100644 --- a/Mathlib/Data/Nat/Cast/Synonym.lean +++ b/Mathlib/Data/Nat/Cast/Synonym.lean @@ -42,20 +42,18 @@ instance [h : AddCommMonoidWithOne α] : AddCommMonoidWithOne αᵒᵈ := theorem toDual_natCast [NatCast α] (n : ℕ) : toDual (n : α) = n := rfl --- See note [no_index around OfNat.ofNat] @[simp] theorem toDual_ofNat [NatCast α] (n : ℕ) [n.AtLeastTwo] : - (toDual (no_index (OfNat.ofNat n : α))) = OfNat.ofNat n := + (toDual (ofNat(n) : α)) = ofNat(n) := rfl @[simp] theorem ofDual_natCast [NatCast α] (n : ℕ) : (ofDual n : α) = n := rfl --- See note [no_index around OfNat.ofNat] @[simp] theorem ofDual_ofNat [NatCast α] (n : ℕ) [n.AtLeastTwo] : - (ofDual (no_index (OfNat.ofNat n : αᵒᵈ))) = OfNat.ofNat n := + (ofDual (ofNat(n) : αᵒᵈ)) = ofNat(n) := rfl /-! ### Lexicographic order -/ diff --git a/Mathlib/Data/Rat/Cast/Defs.lean b/Mathlib/Data/Rat/Cast/Defs.lean index 4b9f06c0c9542..a5fc916543c00 100644 --- a/Mathlib/Data/Rat/Cast/Defs.lean +++ b/Mathlib/Data/Rat/Cast/Defs.lean @@ -33,9 +33,8 @@ variable [DivisionSemiring α] {q r : ℚ≥0} @[simp, norm_cast] lemma cast_natCast (n : ℕ) : ((n : ℚ≥0) : α) = n := by simp [cast_def] --- See note [no_index around OfNat.ofNat] @[simp, norm_cast] lemma cast_ofNat (n : ℕ) [n.AtLeastTwo] : - no_index (OfNat.ofNat n : ℚ≥0) = (OfNat.ofNat n : α) := cast_natCast _ + (ofNat(n) : ℚ≥0) = (ofNat(n) : α) := cast_natCast _ @[simp, norm_cast] lemma cast_zero : ((0 : ℚ≥0) : α) = 0 := (cast_natCast _).trans Nat.cast_zero @[simp, norm_cast] lemma cast_one : ((1 : ℚ≥0) : α) = 1 := (cast_natCast _).trans Nat.cast_one @@ -118,9 +117,8 @@ theorem cast_natCast (n : ℕ) : ((n : ℚ) : α) = n := by @[deprecated (since := "2024-03-21")] alias cast_coe_int := cast_intCast @[deprecated (since := "2024-03-21")] alias cast_coe_nat := cast_natCast --- See note [no_index around OfNat.ofNat] @[simp, norm_cast] lemma cast_ofNat (n : ℕ) [n.AtLeastTwo] : - ((no_index (OfNat.ofNat n : ℚ)) : α) = (OfNat.ofNat n : α) := by + ((ofNat(n) : ℚ) : α) = (ofNat(n) : α) := by simp [cast_def] @[simp, norm_cast] diff --git a/Mathlib/Data/Rat/Defs.lean b/Mathlib/Data/Rat/Defs.lean index 4c5c728579ffa..cb77cd4efc0c3 100644 --- a/Mathlib/Data/Rat/Defs.lean +++ b/Mathlib/Data/Rat/Defs.lean @@ -51,9 +51,8 @@ theorem ofInt_eq_cast (n : ℤ) : ofInt n = Int.cast n := rfl -- TODO: Replace `Rat.ofNat_num`/`Rat.ofNat_den` in Batteries --- See note [no_index around OfNat.ofNat] -@[simp] lemma num_ofNat (n : ℕ) : num (no_index (OfNat.ofNat n)) = OfNat.ofNat n := rfl -@[simp] lemma den_ofNat (n : ℕ) : den (no_index (OfNat.ofNat n)) = 1 := rfl +@[simp] lemma num_ofNat (n : ℕ) : num ofNat(n) = ofNat(n) := rfl +@[simp] lemma den_ofNat (n : ℕ) : den ofNat(n) = 1 := rfl @[simp, norm_cast] lemma num_natCast (n : ℕ) : num n = n := rfl @@ -386,11 +385,11 @@ theorem den_neg_eq_den (q : ℚ) : (-q).den = q.den := theorem num_neg_eq_neg_num (q : ℚ) : (-q).num = -q.num := rfl -@[simp] +-- Not `@[simp]` as `num_ofNat` is stronger. theorem num_zero : Rat.num 0 = 0 := rfl -@[simp] +-- Not `@[simp]` as `den_ofNat` is stronger. theorem den_zero : Rat.den 0 = 1 := rfl @@ -399,7 +398,7 @@ lemma zero_of_num_zero {q : ℚ} (hq : q.num = 0) : q = 0 := by simpa [hq] using theorem zero_iff_num_zero {q : ℚ} : q = 0 ↔ q.num = 0 := ⟨fun _ => by simp [*], zero_of_num_zero⟩ -@[simp] +-- `Not `@[simp]` as `num_ofNat` is stronger. theorem num_one : (1 : ℚ).num = 1 := rfl diff --git a/Mathlib/Lean/Expr/Basic.lean b/Mathlib/Lean/Expr/Basic.lean index 57caa17facbf9..4606e2f6b5692 100644 --- a/Mathlib/Lean/Expr/Basic.lean +++ b/Mathlib/Lean/Expr/Basic.lean @@ -262,6 +262,7 @@ section recognizers partial def numeral? (e : Expr) : Option Nat := if let some n := e.rawNatLit? then n else + let e := e.consumeMData -- `OfNat` numerals may have `no_index` around them from `ofNat()` let f := e.getAppFn if !f.isConst then none else diff --git a/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean b/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean index 3703485af2737..9e0eff961d50c 100644 --- a/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean +++ b/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean @@ -61,7 +61,12 @@ theorem eventually_residual_liouville : ∀ᶠ x in residual ℝ, Liouville x := refine fun n => ⟨r.num * 2, r.den * 2, ?_, ?_⟩ · have := r.pos; omega · convert @mem_ball_self ℝ _ (r : ℝ) _ _ - · push_cast; norm_cast; simp [Rat.divInt_mul_right (two_ne_zero), Rat.mkRat_self] + · push_cast + -- Workaround for https://github.com/leanprover/lean4/pull/6438; this eliminates an + -- `Expr.mdata` that would cause `norm_cast` to skip a numeral. + rw [Eq.refl (2 : ℝ)] + norm_cast + simp [Rat.divInt_mul_right (two_ne_zero), Rat.mkRat_self] · refine one_div_pos.2 (pow_pos (Int.cast_pos.2 ?_) _) exact mul_pos (Int.natCast_pos.2 r.pos) zero_lt_two diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index b9bfe92da053b..7938e6809683b 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -188,6 +188,7 @@ import Mathlib.Tactic.NormNum.Prime import Mathlib.Tactic.NormNum.Result import Mathlib.Tactic.NthRewrite import Mathlib.Tactic.Observe +import Mathlib.Tactic.OfNat import Mathlib.Tactic.PPWithUniv import Mathlib.Tactic.Peel import Mathlib.Tactic.Polyrith diff --git a/Mathlib/Tactic/Common.lean b/Mathlib/Tactic/Common.lean index 94a9e57234ef3..b59d768f32d0c 100644 --- a/Mathlib/Tactic/Common.lean +++ b/Mathlib/Tactic/Common.lean @@ -72,6 +72,7 @@ import Mathlib.Tactic.MkIffOfInductiveProp -- import Mathlib.Tactic.NormNum.Basic import Mathlib.Tactic.NthRewrite import Mathlib.Tactic.Observe +import Mathlib.Tactic.OfNat -- `positivity` imports `Data.Nat.Factorial.Basic`, but hopefully this can be rearranged. -- import Mathlib.Tactic.Positivity import Mathlib.Tactic.ProjectionNotation diff --git a/Mathlib/Tactic/OfNat.lean b/Mathlib/Tactic/OfNat.lean new file mode 100644 index 0000000000000..fb3dbeebb64bd --- /dev/null +++ b/Mathlib/Tactic/OfNat.lean @@ -0,0 +1,18 @@ +/- +Copyright (c) 2024 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import Mathlib.Init + +/-! # The `ofNat()` macro -/ + +/-- +This macro is a shorthand for `OfNat.ofNat` combined with `no_index`. + +When writing lemmas about `OfNat.ofNat`, the term needs to be wrapped +in `no_index` so as not to confuse `simp`, as `no_index (OfNat.ofNat n)`. + +Some discussion is [on Zulip here](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/.E2.9C.94.20Polynomial.2Ecoeff.20example/near/395438147). +-/ +macro "ofNat(" n:term ")" : term => `(no_index (OfNat.ofNat $n)) diff --git a/Mathlib/Tactic/Rify.lean b/Mathlib/Tactic/Rify.lean index b8538585d92da..77801b38118d5 100644 --- a/Mathlib/Tactic/Rify.lean +++ b/Mathlib/Tactic/Rify.lean @@ -82,8 +82,7 @@ macro_rules @[deprecated (since := "2024-04-17")] alias rat_cast_ne := ratCast_ne --- See note [no_index around OfNat.ofNat] @[rify_simps] lemma ofNat_rat_real (a : ℕ) [a.AtLeastTwo] : - ((no_index (OfNat.ofNat a : ℚ)) : ℝ) = (OfNat.ofNat a : ℝ) := rfl + ((ofNat(a) : ℚ) : ℝ) = (ofNat(a) : ℝ) := rfl end Mathlib.Tactic.Rify diff --git a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean index ab3605b44b21f..c32fd5163a6e0 100644 --- a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean @@ -1220,10 +1220,9 @@ instance : NatCast (α →ᵇ R) := @[simp, norm_cast] theorem coe_natCast (n : ℕ) : ((n : α →ᵇ R) : α → R) = n := rfl --- See note [no_index around OfNat.ofNat] @[simp, norm_cast] theorem coe_ofNat (n : ℕ) [n.AtLeastTwo] : - ((no_index (OfNat.ofNat n) : α →ᵇ R) : α → R) = OfNat.ofNat n := + ((ofNat(n) : α →ᵇ R) : α → R) = ofNat(n) := rfl instance : IntCast (α →ᵇ R) := diff --git a/scripts/noshake.json b/scripts/noshake.json index 22f216add9304..a45672285115c 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -450,6 +450,7 @@ "Mathlib.Algebra.Group.Units.Basic": ["Mathlib.Tactic.Subsingleton"], "Mathlib.Algebra.Group.Units": ["Mathlib.Tactic.Subsingleton"], "Mathlib.Algebra.Group.Pi.Basic": ["Batteries.Tactic.Classical"], + "Mathlib.Algebra.Group.Defs": ["Mathlib.Tactic.OfNat"], "Mathlib.Algebra.GeomSum": ["Mathlib.Algebra.Order.BigOperators.Ring.Finset"], "Mathlib.Algebra.Category.Ring.Basic": ["Mathlib.CategoryTheory.ConcreteCategory.ReflectsIso"], From 12bb04fb4d5e3ce98e7e76da8d509453d1b87e2c Mon Sep 17 00:00:00 2001 From: kvanvels Date: Wed, 1 Jan 2025 21:51:08 +0000 Subject: [PATCH 056/189] feat: add `Set.preimage_image_univ` and associated `GaloisConnection` result (#20346) this function shows that the preimage of the image of a universal set is the universal set. Since preimage and image form a Galois connection, this is a special case of a more general theorem that I added there. I chose not to use the general theorem in the special theorem since doing so would require either putting the preimage_image_univ function in an unnatural spot or solving a circular import problem. Co-authored-by: kvanvels <101572210+kvanvels@users.noreply.github.com> --- Mathlib/Data/Set/Image.lean | 3 +++ Mathlib/Order/GaloisConnection.lean | 6 ++++++ 2 files changed, 9 insertions(+) diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 16c6dcc43e64f..2df34e18001a6 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -421,6 +421,9 @@ theorem image_preimage_subset (f : α → β) (s : Set β) : f '' (f ⁻¹' s) theorem subset_preimage_image (f : α → β) (s : Set α) : s ⊆ f ⁻¹' (f '' s) := fun _ => mem_image_of_mem f +theorem preimage_image_univ {f : α → β} : f ⁻¹' (f '' univ) = univ := + Subset.antisymm (fun _ _ => trivial) (subset_preimage_image f univ) + @[simp] theorem preimage_image_eq {f : α → β} (s : Set α) (h : Injective f) : f ⁻¹' (f '' s) = s := Subset.antisymm (fun _ ⟨_, hy, e⟩ => h e ▸ hy) (subset_preimage_image f s) diff --git a/Mathlib/Order/GaloisConnection.lean b/Mathlib/Order/GaloisConnection.lean index c3915d09e0a45..c5ad8900514d8 100644 --- a/Mathlib/Order/GaloisConnection.lean +++ b/Mathlib/Order/GaloisConnection.lean @@ -204,6 +204,9 @@ theorem u_eq_top {l : α → β} {u : β → α} (gc : GaloisConnection l u) {x} theorem u_top [OrderTop β] {l : α → β} {u : β → α} (gc : GaloisConnection l u) : u ⊤ = ⊤ := gc.u_eq_top.2 le_top +theorem u_l_top {l : α → β} {u : β → α} (gc : GaloisConnection l u) : u (l ⊤) = ⊤ := + gc.u_eq_top.mpr le_rfl + end OrderTop section OrderBot @@ -216,6 +219,9 @@ theorem l_eq_bot {l : α → β} {u : β → α} (gc : GaloisConnection l u) {x} theorem l_bot [OrderBot α] {l : α → β} {u : β → α} (gc : GaloisConnection l u) : l ⊥ = ⊥ := gc.dual.u_top +theorem l_u_bot {l : α → β} {u : β → α} (gc : GaloisConnection l u) : l (u ⊥) = ⊥ := + gc.l_eq_bot.mpr le_rfl + end OrderBot section SemilatticeSup From ec10d9ba44225d66e790787d0f5b19a14830337e Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 1 Jan 2025 22:55:48 +0000 Subject: [PATCH 057/189] chore: split `Lean/Meta/RefinedDiscrTree` (#20381) This split is intended to roughly match #11968, in order to reduce the diff there. I will take ownership of fixing any resulting merge conflicts. The declarations have been copied unchanged between files, with the exception of a single removed `private` on `DTExpr.eqv` that prevented this split. Refactors are left to #11968. The namespace of all the declarations has changed from `Mathlib.Meta.FunProp.RefinedDiscrTree` to `Lean.Meta.RefinedDiscrTree`, to match the filename. --- Mathlib.lean | 4 + Mathlib/Lean/Meta/RefinedDiscrTree.lean | 939 +----------------- Mathlib/Lean/Meta/RefinedDiscrTree/Basic.lean | 179 ++++ .../Lean/Meta/RefinedDiscrTree/Encode.lean | 508 ++++++++++ .../Lean/Meta/RefinedDiscrTree/Lookup.lean | 187 ++++ Mathlib/Lean/Meta/RefinedDiscrTree/Pi.lean | 115 +++ Mathlib/Tactic/FunProp/Theorems.lean | 1 + scripts/noshake.json | 2 +- 8 files changed, 1001 insertions(+), 934 deletions(-) create mode 100644 Mathlib/Lean/Meta/RefinedDiscrTree/Basic.lean create mode 100644 Mathlib/Lean/Meta/RefinedDiscrTree/Encode.lean create mode 100644 Mathlib/Lean/Meta/RefinedDiscrTree/Lookup.lean create mode 100644 Mathlib/Lean/Meta/RefinedDiscrTree/Pi.lean diff --git a/Mathlib.lean b/Mathlib.lean index 2de875ff969de..6ff4958551d47 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3322,6 +3322,10 @@ import Mathlib.Lean.Meta.CongrTheorems import Mathlib.Lean.Meta.DiscrTree import Mathlib.Lean.Meta.KAbstractPositions import Mathlib.Lean.Meta.RefinedDiscrTree +import Mathlib.Lean.Meta.RefinedDiscrTree.Basic +import Mathlib.Lean.Meta.RefinedDiscrTree.Encode +import Mathlib.Lean.Meta.RefinedDiscrTree.Lookup +import Mathlib.Lean.Meta.RefinedDiscrTree.Pi import Mathlib.Lean.Meta.Simp import Mathlib.Lean.Name import Mathlib.Lean.PrettyPrinter.Delaborator diff --git a/Mathlib/Lean/Meta/RefinedDiscrTree.lean b/Mathlib/Lean/Meta/RefinedDiscrTree.lean index 439560f236c24..0ec2891be7c59 100644 --- a/Mathlib/Lean/Meta/RefinedDiscrTree.lean +++ b/Mathlib/Lean/Meta/RefinedDiscrTree.lean @@ -3,8 +3,9 @@ Copyright (c) 2023 J. W. Gerbscheid. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: J. W. Gerbscheid -/ -import Mathlib.Tactic.FunProp.StateList -import Mathlib.Algebra.Group.Pi.Basic +import Mathlib.Lean.Meta.RefinedDiscrTree.Pi +import Mathlib.Lean.Meta.RefinedDiscrTree.Basic +import Mathlib.Lean.Meta.RefinedDiscrTree.Encode /-! We define discrimination trees for the purpose of unifying local expressions with library results. @@ -105,761 +106,8 @@ TODO: open Lean Meta -namespace Mathlib.Meta.FunProp.RefinedDiscrTree - -/-! ## Definitions -/ - -/-- Discrimination tree key. -/ -inductive Key where - /-- A metavariable. This key matches with anything. It stores an index. -/ - | star : Nat → Key - /-- An opaque variable. This key only matches with itself or `Key.star`. -/ - | opaque : Key - /-- A constant. It stores the name and the arity. -/ - | const : Name → Nat → Key - /-- A free variable. It stores the `FVarId` and the arity. -/ - | fvar : FVarId → Nat → Key - /-- A bound variable, from a lambda or forall binder. - It stores the De Bruijn index and the arity. -/ - | bvar : Nat → Nat → Key - /-- A literal. -/ - | lit : Literal → Key - /-- A sort. Universe levels are ignored. -/ - | sort : Key - /-- A lambda function. -/ - | lam : Key - /-- A dependent arrow. -/ - | forall : Key - /-- A projection. It stores the structure name, the projection index and the arity. -/ - | proj : Name → Nat → Nat → Key - deriving Inhabited, BEq, Repr - -private nonrec def Key.hash : Key → UInt64 - | .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 - | .sort => 2411 - | .lam => 4742 - | .«forall» => 9752 - | .proj s i a => mixHash (hash a) <| mixHash (hash s) (hash i) - -instance : Hashable Key := ⟨Key.hash⟩ - -/-- Constructor index used for ordering `Key`. -Note that the index of the star pattern is 0, so that when looking up in a `Trie`, -we can look at the start of the sorted array for all `.star` patterns. -/ -def Key.ctorIdx : Key → Nat - | .star .. => 0 - | .opaque .. => 1 - | .const .. => 2 - | .fvar .. => 3 - | .bvar .. => 4 - | .lit .. => 5 - | .sort => 6 - | .lam => 7 - | .forall => 8 - | .proj .. => 9 - -/-- The order on `Key` used in the `RefinedDiscrTree`. -/ -private def Key.lt : Key → Key → Bool - | .star i₁, .star i₂ => i₁ < i₂ - | .const n₁ a₁, .const n₂ a₂ => Name.quickLt n₁ n₂ || (n₁ == n₂ && a₁ < a₂) - | .fvar f₁ a₁, .fvar f₂ a₂ => Name.quickLt f₁.name f₂.name || (f₁ == f₂ && a₁ < a₂) - | .bvar i₁ a₁, .bvar i₂ a₂ => i₁ < i₂ || (i₁ == i₂ && a₁ < a₂) - | .lit v₁, .lit v₂ => v₁ < v₂ - | .proj s₁ i₁ a₁, .proj s₂ i₂ a₂ => Name.quickLt s₁ s₂ || - (s₁ == s₂ && (i₁ < i₂ || (i₁ == i₂ && a₁ < a₂))) - | k₁, k₂ => k₁.ctorIdx < k₂.ctorIdx - -instance : LT Key := ⟨fun a b => Key.lt a b⟩ -instance (a b : Key) : Decidable (a < b) := inferInstanceAs (Decidable (Key.lt a b)) - -private def Key.format : Key → Format - | .star i => "*" ++ Std.format i - | .opaque => "◾" - | .const k a => "⟨" ++ Std.format k ++ ", " ++ Std.format a ++ "⟩" - | .fvar k a => "⟨" ++ Std.format k.name ++ ", " ++ Std.format a ++ "⟩" - | .lit (Literal.natVal v) => Std.format v - | .lit (Literal.strVal v) => repr v - | .sort => "sort" - | .bvar i a => "⟨" ++ "#" ++ Std.format i ++ ", " ++ Std.format a ++ "⟩" - | .lam => "λ" - | .forall => "∀" - | .proj s i a => "⟨" ++ Std.format s ++"."++ Std.format i ++", "++ Std.format a ++ "⟩" - -instance : ToFormat Key := ⟨Key.format⟩ - -/-- Return the number of arguments that the `Key` takes. -/ -def Key.arity : Key → Nat - | .const _ a => a - | .fvar _ a => a - | .bvar _ a => a - | .lam => 1 - | .forall => 2 - | .proj _ _ a => 1 + a - | _ => 0 - -variable {α : Type} -/-- Discrimination tree trie. See `RefinedDiscrTree`. -/ -inductive Trie (α : Type) where - /-- Map from `Key` to `Trie`. Children is an `Array` of size at least 2, - sorted in increasing order using `Key.lt`. -/ - | node (children : Array (Key × Trie α)) - /-- Sequence of nodes with only one child. `keys` is an `Array` of size at least 1. -/ - | path (keys : Array Key) (child : Trie α) - /-- Leaf of the Trie. `values` is an `Array` of size at least 1. -/ - | values (vs : Array α) -instance : Inhabited (Trie α) := ⟨.node #[]⟩ - -/-- `Trie.path` constructor that only inserts the path if it is non-empty. -/ -def Trie.mkPath (keys : Array Key) (child : Trie α) := - if keys.isEmpty then child else Trie.path keys child - -/-- `Trie` constructor for a single value, taking the keys starting at index `i`. -/ -def Trie.singleton (keys : Array Key) (value : α) (i : Nat) : Trie α := - mkPath keys[i:] (values #[value]) - -/-- `Trie.node` constructor for combining two `Key`, `Trie α` pairs. -/ -def Trie.mkNode2 (k1 : Key) (t1 : Trie α) (k2 : Key) (t2 : Trie α) : Trie α := - if k1 < k2 then - .node #[(k1, t1), (k2, t2)] - else - .node #[(k2, t2), (k1, t1)] - -/-- Return the values from a `Trie α`, assuming that it is a leaf -/ -def Trie.values! : Trie α → Array α - | .values vs => vs - | _ => panic! "expected .values constructor" - -/-- Return the children of a `Trie α`, assuming that it is not a leaf. -The result is sorted by the `Key`'s -/ -def Trie.children! : Trie α → Array (Key × Trie α) - | .node cs => cs - | .path ks c => #[(ks[0]!, mkPath ks[1:] c)] - | .values _ => panic! "did not expect .values constructor" - -private partial def Trie.format [ToFormat α] : Trie α → Format - | .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 - | .path ks c => Format.sbracket (Format.joinSep ks.toList (", ")) - ++ " => " ++ Format.line ++ format c -where - prepend (k : Key) (t : Trie α) : Trie α := match t with - | .path ks c => .path (#[k] ++ ks) c - | t => .path #[k] t -instance [ToFormat α] : ToFormat (Trie α) := ⟨Trie.format⟩ - - -/-- Discrimination tree. It is an index from expressions to values of type `α`. -/ -structure _root_.Mathlib.Meta.FunProp.RefinedDiscrTree (α : Type) where - /-- The underlying `PersistentHashMap` of a `RefinedDiscrTree`. -/ - root : PersistentHashMap Key (Trie α) := {} -instance : Inhabited (RefinedDiscrTree α) := ⟨{}⟩ - -private partial def format [ToFormat α] (d : RefinedDiscrTree α) : Format := - let (_, r) := d.root.foldl - (fun (p : Bool × Format) k c => - (false, - p.2 ++ (if p.1 then Format.nil else Format.line) ++ - Format.paren (Std.format k ++ " => " ++ Std.format c))) - (true, Format.nil) - Format.group r - -instance [ToFormat α] : ToFormat (RefinedDiscrTree α) := ⟨format⟩ - - -/-- `DTExpr` is a simplified form of `Expr`. -It is the intermediate step for converting from `Expr` to `Array Key`. -/ -inductive DTExpr where - /-- A metavariable. It optionally stores an `MVarId`. -/ - | star : Option MVarId → DTExpr - /-- An opaque variable or a let-expression in the case `WhnfCoreConfig.zeta := false`. -/ - | opaque : DTExpr - /-- A constant. It stores the name and the arguments. -/ - | const : Name → Array DTExpr → DTExpr - /-- A free variable. It stores the `FVarId` and the arguments -/ - | fvar : FVarId → Array DTExpr → DTExpr - /-- A bound variable. It stores the De Bruijn index and the arguments -/ - | bvar : Nat → Array DTExpr → DTExpr - /-- A literal. -/ - | lit : Literal → DTExpr - /-- A sort. -/ - | sort : DTExpr - /-- A lambda function. It stores the body. -/ - | lam : DTExpr → DTExpr - /-- A dependent arrow. It stores the domain and body. -/ - | forall : DTExpr → DTExpr → DTExpr - /-- A projection. It stores the structure name, projection index, struct body and arguments. -/ - | proj : Name → Nat → DTExpr → Array DTExpr → DTExpr -deriving Inhabited, BEq, Repr - -private partial def DTExpr.format : DTExpr → Format - | .star _ => "*" - | .opaque => "◾" - | .const n as => Std.format n ++ formatArgs as - | .fvar n as => Std.format n.name ++ formatArgs as - | .bvar i as => "#" ++ Std.format i ++ formatArgs as - | .lit (Literal.natVal v) => Std.format v - | .lit (Literal.strVal v) => repr v - | .sort => "Sort" - | .lam b => "λ " ++ DTExpr.format b - | .forall d b => DTExpr.format d ++ " → " ++ DTExpr.format b - | .proj _ i a as => DTExpr.format a ++ "." ++ Std.format i ++ formatArgs as -where - formatArgs (as : Array DTExpr) := - if as.isEmpty - then .nil - else " " ++ Format.paren (@Format.joinSep _ ⟨DTExpr.format⟩ as.toList ", ") - -instance : ToFormat DTExpr := ⟨DTExpr.format⟩ - -/-- Return the size of the `DTExpr`. This is used for calculating the matching score when two -expressions are equal. -The score is not incremented at a lambda, which is so that the expressions -`∀ x, p[x]` and `∃ x, p[x]` get the same size. -/ -partial def DTExpr.size : DTExpr → Nat -| .const _ args -| .fvar _ args -| .bvar _ args => args.foldl (init := 1) (· + ·.size) -| .lam b => b.size -| .forall d b => 1 + d.size + b.size -| _ => 1 - -private def DTExpr.eqv (a b : DTExpr) : Bool := - (go a b).run' {} -where - 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₂ - | .fvar n₁ as₁ , .fvar n₂ as₂ => pure (n₁ == n₂) <&&> goArray as₁ as₂ - | .bvar i₁ as₁ , .bvar i₂ as₂ => pure (i₁ == i₂) <&&> goArray as₁ as₂ - | .lit li₁ , .lit li₂ => pure (li₁ == li₂) - | .sort , .sort => pure true - | .lam b₁ , .lam b₂ => go b₁ b₂ - | .forall d₁ b₁ , .forall d₂ b₂ => go d₁ d₂ <&&> go b₁ b₂ - | .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[id₁]? with - | some id => (id == id₂, map) - | none => (true, map.insert id₁ id₂) - | _ , _ => return false - - 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 - return false - return true - else - return false - -/-! ## Encoding an Expr -/ - -/-- This state is used to turn the indexing by `MVarId` and `FVarId` in `DTExpr` into -indexing by `Nat` in `Key`. -/ -private structure Flatten.State where - stars : Array MVarId := #[] - -private def getStar (mvarId? : Option MVarId) : StateM Flatten.State Nat := - modifyGet fun s => - match mvarId? with - | some mvarId => match s.stars.findIdx? (· == mvarId) with - | some idx => (idx, s) - | none => (s.stars.size, { s with stars := s.stars.push mvarId }) - | none => (s.stars.size, { s with stars := s.stars.push ⟨.anonymous⟩ }) - -private partial def DTExpr.flattenAux (todo : Array Key) : DTExpr → StateM Flatten.State (Array Key) - | .star i => return todo.push (.star (← getStar i)) - | .opaque => return todo.push .opaque - | .const n as => as.foldlM flattenAux (todo.push (.const n as.size)) - | .fvar f as => as.foldlM flattenAux (todo.push (.fvar f as.size)) - | .bvar i as => as.foldlM flattenAux (todo.push (.bvar i as.size)) - | .lit l => return todo.push (.lit l) - | .sort => return todo.push .sort - | .lam b => flattenAux (todo.push .lam) b - | .«forall» d b => do flattenAux (← flattenAux (todo.push .forall) d) b - | .proj n i e as => do as.foldlM flattenAux (← flattenAux (todo.push (.proj n i as.size)) e) - -/-- Given a `DTExpr`, return the linearized encoding in terms of `Key`, -which is used for `RefinedDiscrTree` indexing. -/ -def DTExpr.flatten (e : DTExpr) (initCapacity := 16) : Array Key := - (DTExpr.flattenAux (.mkEmpty initCapacity) e).run' {} - - - -/-- Return true if `e` is one of the following -- A nat literal (numeral) -- `Nat.zero` -- `Nat.succ x` where `isNumeral x` -- `OfNat.ofNat _ x _` where `isNumeral x` -/ -private partial def isNumeral (e : Expr) : Bool := - if e.isRawNatLit then true - else - let f := e.getAppFn - if !f.isConst then false - else - let fName := f.constName! - if fName == ``Nat.succ && e.getAppNumArgs == 1 then isNumeral e.appArg! - else if fName == ``OfNat.ofNat && e.getAppNumArgs == 3 then isNumeral (e.getArg! 1) - else if fName == ``Nat.zero && e.getAppNumArgs == 0 then true - else false - -/-- Return `some n` if `e` is definitionally equal to the natural number `n`. -/ -private partial def toNatLit? (e : Expr) : Option Literal := - if isNumeral e then - if let some n := loop e then - some (.natVal n) - else - none - else - none -where - loop (e : Expr) : Option Nat := do - let f := e.getAppFn - match f with - | .lit (.natVal n) => return n - | .const fName .. => - if fName == ``Nat.succ && e.getAppNumArgs == 1 then - let r ← loop e.appArg! - return r+1 - else if fName == ``OfNat.ofNat && e.getAppNumArgs == 3 then - loop (e.getArg! 1) - else if fName == ``Nat.zero && e.getAppNumArgs == 0 then - return 0 - else - failure - | _ => failure - -/-- Reduction procedure for the `RefinedDiscrTree` indexing. -/ -partial def reduce (e : Expr) : MetaM Expr := do - let e ← whnfCore e - match (← unfoldDefinition? e) with - | some e => reduce e - | none => match e.etaExpandedStrict? with - | some e => reduce e - | none => return e - -/-- Repeatedly apply reduce while stripping lambda binders and introducing their variables -/ -@[specialize] -partial def lambdaTelescopeReduce {m} [Monad m] [MonadLiftT MetaM m] [MonadControlT MetaM m] - [Nonempty α] (e : Expr) (fvars : List FVarId) - (k : Expr → List FVarId → m α) : m α := do - match ← reduce e with - | .lam n d b bi => - withLocalDecl n bi d fun fvar => - lambdaTelescopeReduce (b.instantiate1 fvar) (fvar.fvarId! :: fvars) k - | e => k e fvars - - - -/-- Check whether the expression is represented by `Key.star`. -/ -def isStar : Expr → Bool - | .mvar .. => true - | .app f _ => isStar f - | _ => false - -/-- Check whether the expression is represented by `Key.star` and has `arg` as an argument. -/ -def isStarWithArg (arg : Expr) : Expr → Bool - | .app f a => if a == arg then isStar f else isStarWithArg arg f - | _ => false - -private partial def DTExpr.hasLooseBVarsAux (i : Nat) : DTExpr → Bool - | .const _ as => as.any (hasLooseBVarsAux i) - | .fvar _ as => as.any (hasLooseBVarsAux i) - | .bvar j as => j ≥ i || as.any (hasLooseBVarsAux i) - | .proj _ _ a as => a.hasLooseBVarsAux i || as.any (hasLooseBVarsAux i) - | .forall d b => d.hasLooseBVarsAux i || b.hasLooseBVarsAux (i+1) - | .lam b => b.hasLooseBVarsAux (i+1) - | _ => false - -/-- Return `true` if `e` contains a loose bound variable. -/ -def DTExpr.hasLooseBVars (e : DTExpr) : Bool := - e.hasLooseBVarsAux 0 - - -namespace MkDTExpr - -private structure Context where - /-- Variables that come from a lambda or forall binder. - The list index gives the De Bruijn index. -/ - bvars : List FVarId := [] - /-- Variables that come from a lambda that has been removed via η-reduction. -/ - forbiddenVars : List FVarId := [] - fvarInContext : FVarId → Bool - -/-- Return for each argument whether it should be ignored. -/ -def getIgnores (fn : Expr) (args : Array Expr) : MetaM (Array Bool) := do - let mut fnType ← inferType fn - let mut result := Array.mkEmpty args.size - let mut j := 0 - for i in [:args.size] do - unless fnType matches .forallE .. do - fnType ← whnfD (fnType.instantiateRevRange j i args) - j := i - let .forallE _ d b bi := fnType | throwError m! "expected function type {indentExpr fnType}" - fnType := b - result := result.push (← isIgnoredArg args[i]! d bi) - return result -where - /-- Return whether the argument should be ignored. -/ - isIgnoredArg (arg domain : Expr) (binderInfo : BinderInfo) : MetaM Bool := do - if domain.isOutParam then - return true - match binderInfo with - | .instImplicit => return true - | .implicit - | .strictImplicit => return !(← isType arg) - | .default => isProof arg - - - -/-- Introduce new lambdas by η-expansion. -/ -@[specialize] -def etaExpand (args : Array Expr) (type : Expr) (lambdas : List FVarId) (goalArity : Nat) - (k : Array Expr → List FVarId → MetaM α) : MetaM α := do - if args.size < goalArity then - withLocalDeclD `_η type fun fvar => - etaExpand (args.push fvar) type (fvar.fvarId! :: lambdas) goalArity k - else - k args lambdas - - -/-- Normalize an application of a heterogeneous binary operator like `HAdd.hAdd`, using: -- `f = fun x => f x` to increase the arity to 6 -- `(f + g) a = f a + g a` to decrease the arity to 6 -- `(fun x => f x + g x) = f + g` to get rid of any lambdas in front -/ -def reduceHBinOpAux (args : Array Expr) (lambdas : List FVarId) (instH instPi : Name) : - OptionT MetaM (Expr × Expr × Expr × List FVarId) := do - let some (mkApp2 (.const instH' _) type inst) := args[3]? | failure - guard (instH == instH') - if args.size ≤ 6 then - etaExpand args type lambdas 6 fun args lambdas => - distributeLambdas lambdas type args[4]! args[5]! - else - /- use that `(f + g) a = f a + g a` -/ - let mut type := type - let mut inst := inst - let mut lhs := args[4]! - let mut rhs := args[5]! - for arg in args[6:] do - let mkApp3 (.const i _) _ f inst' := inst | return (type, lhs, rhs, lambdas) - unless i == instPi do return (type, lhs, rhs, lambdas) - type := .app f arg - inst := inst' - lhs := .app lhs arg - rhs := .app rhs arg - distributeLambdas lambdas type lhs rhs -where - /-- use that `(fun x => f x + g x) = f + g` -/ - distributeLambdas (lambdas : List FVarId) (type lhs rhs : Expr) : - MetaM (Expr × Expr × Expr × List FVarId) := match lambdas with - | fvarId :: lambdas => do - let decl ← fvarId.getDecl - let type := .forallE decl.userName decl.type (type.abstract #[.fvar fvarId]) decl.binderInfo - let lhs := .lam decl.userName decl.type (lhs.abstract #[.fvar fvarId]) decl.binderInfo - let rhs := .lam decl.userName decl.type (rhs.abstract #[.fvar fvarId]) decl.binderInfo - distributeLambdas lambdas type lhs rhs - | [] => return (type, lhs, rhs, []) - -/-- Normalize an application if the head is `+`, `*`, `-` or `/`. -Optionally return the `(type, lhs, rhs, lambdas)`. -/ -@[inline] def reduceHBinOp (n : Name) (args : Array Expr) (lambdas : List FVarId) : - MetaM (Option (Expr × Expr × Expr × List FVarId)) := - match n with - | ``HAdd.hAdd => reduceHBinOpAux args lambdas ``instHAdd ``Pi.instAdd - | ``HMul.hMul => reduceHBinOpAux args lambdas ``instHMul ``Pi.instMul - | ``HSub.hSub => reduceHBinOpAux args lambdas ``instHSub ``Pi.instSub - | ``HDiv.hDiv => reduceHBinOpAux args lambdas ``instHDiv ``Pi.instDiv - | _ => return none - -/-- Normalize an application of a unary operator like `Inv.inv`, using: -- `f⁻¹ a = (f a)⁻¹` to decrease the arity to 3 -- `(fun x => (f a)⁻¹) = f⁻¹` to get rid of any lambdas in front -/ -def reduceUnOpAux (args : Array Expr) (lambdas : List FVarId) (instPi : Name) : - OptionT MetaM (Expr × Expr × List FVarId) := do - guard (args.size ≥ 3) - let mut type := args[0]! - let mut inst := args[1]! - let mut arg := args[2]! - if args.size == 3 then - distributeLambdas lambdas type arg - else - /- use that `f⁻¹ a = (f a)⁻¹` -/ - for arg' in args[3:] do - let mkApp3 (.const i _) _ f inst' := inst | return (type, arg, lambdas) - unless i == instPi do return (type, arg, lambdas) - type := .app f arg' - inst := inst' - arg := .app arg arg' - distributeLambdas lambdas type arg -where - /-- use that `(fun x => (f x)⁻¹) = f⁻¹` -/ - distributeLambdas (lambdas : List FVarId) (type arg : Expr) : - MetaM (Expr × Expr × List FVarId) := match lambdas with - | fvarId :: lambdas => do - let decl ← fvarId.getDecl - let type := .forallE decl.userName decl.type (type.abstract #[.fvar fvarId]) decl.binderInfo - let arg := .lam decl.userName decl.type (arg.abstract #[.fvar fvarId]) decl.binderInfo - distributeLambdas lambdas type arg - | [] => return (type, arg, []) - -/-- Normalize an application if the head is `⁻¹` or `-`. -Optionally return the `(type, arg, lambdas)`. -/ -@[inline] def reduceUnOp (n : Name) (args : Array Expr) (lambdas : List FVarId) : - MetaM (Option (Expr × Expr × List FVarId)) := - match n with - | ``Neg.neg => reduceUnOpAux args lambdas ``Pi.instNeg - | ``Inv.inv => reduceUnOpAux args lambdas ``Pi.instInv - | _ => return none - -@[specialize] -private def withLams {m} [Monad m] [MonadWithReader Context m] - (lambdas : List FVarId) (k : m DTExpr) : m DTExpr := - if lambdas.isEmpty then - k - else do - let e ← withReader (fun c => { c with bvars := lambdas ++ c.bvars }) k - return lambdas.foldl (fun _ => ·.lam) e - - -/-- Return the encoding of `e` as a `DTExpr`. -If `root = false`, then `e` is a strict sub expression of the original expression. -/ -partial def mkDTExprAux (e : Expr) (root : Bool) : ReaderT Context MetaM DTExpr := do - lambdaTelescopeReduce e [] fun e lambdas => - e.withApp fun fn args => do - - let argDTExpr (arg : Expr) (ignore : Bool) : ReaderT Context MetaM DTExpr := - if ignore then pure (.star none) else mkDTExprAux arg false - - let argDTExprs : ReaderT Context MetaM (Array DTExpr) := do - let ignores ← getIgnores fn args - args.mapIdxM fun i arg => - argDTExpr arg ignores[i]! - - match fn with - | .const n _ => - unless root do - if let some (type, lhs, rhs, lambdas') ← reduceHBinOp n args lambdas then - return ← withLams lambdas' do - let type ← mkDTExprAux type false - let lhs ← mkDTExprAux lhs false - let rhs ← mkDTExprAux rhs false - return .const n #[type, type, .star none, .star none, lhs, rhs] - - if let some (type, arg, lambdas') ← reduceUnOp n e.getAppArgs lambdas then - return ← withLams lambdas' do - let type ← mkDTExprAux type false - let arg ← mkDTExprAux arg false - return .const n #[type, .star none, arg] - - /- since `(fun _ => 0) = 0` and `(fun _ => 1) = 1`, - we don't index lambdas before literals -/ - if let some v := toNatLit? e then - return .lit v - withLams lambdas do - return .const n (← argDTExprs) - | .proj s i a => - withLams lambdas do - let a ← argDTExpr a (isClass (← getEnv) s) - return .proj s i a (← argDTExprs) - | .fvar fvarId => - /- we index `fun x => x` as `id` when not at the root -/ - if let fvarId' :: lambdas' := lambdas then - if fvarId' == fvarId && args.isEmpty && !root then - return ← withLams lambdas' do - let type ← mkDTExprAux (← fvarId.getType) false - return .const ``id #[type] - withLams lambdas do - if let some idx := (← read).bvars.findIdx? (· == fvarId) then - return .bvar idx (← argDTExprs) - if (← read).fvarInContext fvarId then - return .fvar fvarId (← argDTExprs) - else - return .opaque - | .mvar mvarId => - /- When the mvarId has arguments, index it with `[*]` instead of `[λ,*]`, - because the star could depend on the bound variables. As a result, - something indexed `[λ,*]` has that the `*` cannot depend on the λ-bound variables -/ - if args.isEmpty then - withLams lambdas do return .star (some mvarId) - else - return .star none - - | .forallE n d b bi => - withLams lambdas do - let d' ← mkDTExprAux d false - let b' ← withLocalDecl n bi d fun fvar => - withReader (fun c => { c with bvars := fvar.fvarId! :: c.bvars }) do - mkDTExprAux (b.instantiate1 fvar) false - return .forall d' b' - | .lit v => withLams lambdas do return .lit v - | .sort _ => withLams lambdas do return .sort - | .letE .. => withLams lambdas do return .opaque - | .lam .. => withLams lambdas do return .opaque - | _ => unreachable! - - -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 -a different number of binders, then the resulting De Bruijn indices are offset. -In practice, getting a `.bvar` in a `DTExpr` is very rare, so we exclude such values from the cache. --/ -instance : MonadCache Expr DTExpr M where - findCached? e := do - let s ← get - return s.find? e - cache e e' := - if e'.hasLooseBVars then - return - else - modify (·.insert e e') - -/-- Return all pairs of body, bound variables that could possibly appear due to η-reduction -/ -@[specialize] -def etaPossibilities (e : Expr) (lambdas : List FVarId) (k : Expr → List FVarId → M α) : M α := - k e lambdas - <|> do - match e, lambdas with - | .app f a, fvarId :: lambdas => - if isStarWithArg (.fvar fvarId) a then - withReader (fun c => { c with forbiddenVars := fvarId :: c.forbiddenVars }) do - etaPossibilities f lambdas k - else - failure - | _, _ => failure - -/-- run `etaPossibilities`, and cache the result if there are multiple possibilities. -/ -@[specialize] -def cacheEtaPossibilities (e original : Expr) (lambdas : List FVarId) - (k : Expr → List FVarId → M DTExpr) : M DTExpr := - match e, lambdas with - | .app _ a, fvarId :: _ => - if isStarWithArg (.fvar fvarId) a then - checkCache original fun _ => - etaPossibilities e lambdas k - else - k e lambdas - | _, _ => k e lambdas - - -/-- Return all encodings of `e` as a `DTExpr`, taking possible η-reductions into account. -If `root = false`, then `e` is a strict sub expression of the original expression. -/ -partial def mkDTExprsAux (original : Expr) (root : Bool) : M DTExpr := do - lambdaTelescopeReduce original [] fun e lambdas => do - - if !root then - if let .const n _ := e.getAppFn then - if let some (type, lhs, rhs, lambdas') ← reduceHBinOp n e.getAppArgs lambdas then - return ← withLams lambdas' do - let type ← mkDTExprsAux type false - let lhs ← mkDTExprsAux lhs false - let rhs ← mkDTExprsAux rhs false - return .const n #[type, type, .star none, .star none, lhs, rhs] - - if let some (type, arg, lambdas') ← reduceUnOp n e.getAppArgs lambdas then - return ← withLams lambdas' do - let type ← mkDTExprsAux type false - let arg ← mkDTExprsAux arg false - return .const n #[type, .star none, arg] - - cacheEtaPossibilities e original lambdas fun e lambdas => - e.withApp fun fn args => do - - let argDTExpr (arg : Expr) (ignore : Bool) : M DTExpr := - if ignore then pure (.star none) else mkDTExprsAux arg false - - let argDTExprs : M (Array DTExpr) := do - let ignores ← getIgnores fn args - args.mapIdxM fun i arg => - argDTExpr arg ignores[i]! - - match fn with - | .const n _ => - unless root do - /- since `(fun _ => 0) = 0` and `(fun _ => 1) = 1`, - we don't index lambdas before nat literals -/ - if let some v := toNatLit? e then - return .lit v - withLams lambdas do - return .const n (← argDTExprs) - | .proj s i a => - withLams lambdas do - let a ← argDTExpr a (isClass (← getEnv) s) - return .proj s i a (← argDTExprs) - | .fvar fvarId => - /- we index `fun x => x` as `id` when not at the root -/ - if let fvarId' :: lambdas' := lambdas then - if fvarId' == fvarId && args.isEmpty && !root then - return ← withLams lambdas' do - let type ← mkDTExprAux (← fvarId.getType) false - return .const ``id #[type] - withLams lambdas do - let c ← read - if let some idx := c.bvars.findIdx? (· == fvarId) then - return .bvar idx (← argDTExprs) - guard !(c.forbiddenVars.contains fvarId) - if c.fvarInContext fvarId then - return .fvar fvarId (← argDTExprs) - else - return .opaque - | .mvar mvarId => - /- When the mvarId has arguments, index it with `[*]` instead of `[λ,*]`, - because the star could depend on the bound variables. As a result, - something indexed `[λ,*]` has that the `*` cannot depend on the λ-bound variables -/ - if args.isEmpty then - withLams lambdas do return .star (some mvarId) - else - return .star none - - | .forallE n d b bi => - withLams lambdas do - let d' ← mkDTExprsAux d false - let b' ← withLocalDecl n bi d fun fvar => - withReader (fun c => { c with bvars := fvar.fvarId! :: c.bvars }) do - mkDTExprsAux (b.instantiate1 fvar) false - return .forall d' b' - | .lit v => withLams lambdas do return .lit v - | .sort _ => withLams lambdas do return .sort - | .letE .. => withLams lambdas do return .opaque - | .lam .. => withLams lambdas do return .opaque - | _ => unreachable! - -end MkDTExpr - -/-- -/ -def DTExpr.isSpecific : DTExpr → Bool - | .star _ - | .const ``Eq #[.star _, .star _, .star _] => false - | _ => true - -/-- Return the encoding of `e` as a `DTExpr`. - -Warning: to account for potential η-reductions of `e`, use `mkDTExprs` instead. - -The argument `fvarInContext` allows you to specify which free variables in `e` will still be -in the context when the `RefinedDiscrTree` is being used for lookup. -It should return true only if the `RefinedDiscrTree` is built and used locally. -/ -def mkDTExpr (e : Expr) - (fvarInContext : FVarId → Bool := fun _ => false) : MetaM DTExpr := - withReducible do (MkDTExpr.mkDTExprAux e true |>.run {fvarInContext}) - -/-- Similar to `mkDTExpr`. -Return all encodings of `e` as a `DTExpr`, taking potential further η-reductions into account. -/ -def mkDTExprs (e : Expr) (onlySpecific : Bool) - (fvarInContext : FVarId → Bool := fun _ => false) : MetaM (List DTExpr) := - withReducible do - let es ← (MkDTExpr.mkDTExprsAux e true).run' {} |>.run {fvarInContext} - return if onlySpecific then es.filter (·.isSpecific) else es - +namespace Lean.Meta.RefinedDiscrTree +variable {α} /-! ## Inserting intro a RefinedDiscrTree -/ @@ -951,181 +199,6 @@ def insertEqn [BEq α] (d : RefinedDiscrTree α) (lhs rhs : Expr) (vLhs vRhs : -/-! ## Matching with a RefinedDiscrTree - -We use a very simple unification algorithm. For all star/metavariable patterns in the -`RefinedDiscrTree` and in the target, we store the assignment, and when it is assigned again, -we check that it is the same assignment. --/ - -namespace GetUnify - -/-- If `k` is a key in `children`, return the corresponding `Trie α`. Otherwise return `none`. -/ -def findKey (children : Array (Key × Trie α)) (k : Key) : Option (Trie α) := - (·.2) <$> children.binSearch (k, default) (fun a b => a.1 < b.1) - -private structure Context where - unify : Bool - -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 : Std.HashMap Nat DTExpr := {} - /-- Metavariable assignments for the `Expr.mvar` in the expression. -/ - mvarAssignments : Std.HashMap MVarId (Array Key) := {} - - -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) (x : M (Trie α)) : - Array (Array α × Nat) := - ((x.run { unify }).run {}).toArray.map (fun (t, s) => (t.values!, s.score)) - -/-- Increment the score by `n`. -/ -private def incrementScore (n : Nat) : M Unit := - modify fun s => { s with score := s.score + n } - -/-- Log a metavariable assignment in the `State`. -/ -private def insertStarAssignment (n : Nat) (e : DTExpr) : M Unit := - modify fun s => { s with starAssignments := s.starAssignments.insert n e } - -/-- Log a metavariable assignment in the `State`. -/ -private def assignMVar (mvarId : MVarId) (e : Array Key) : M Unit := do - let { mvarAssignments, .. } ← get - match mvarAssignments[mvarId]? with - | some e' => guard (e == e') - | none => - modify fun s => { s with mvarAssignments := s.mvarAssignments.insert mvarId e } - -/-- Return the possible `Trie α` that match with `n` metavariable. -/ -partial def skipEntries (t : Trie α) (skipped : Array Key) : Nat → M (Array Key × Trie α) - | 0 => pure (skipped, t) - | skip+1 => - t.children!.foldr (init := failure) fun (k, c) x => - (skipEntries c (skipped.push k) (skip + k.arity)) <|> x -/-- Return the possible `Trie α` that match with anything. -We add 1 to the matching score when the key is `.opaque`, -since this pattern is "harder" to match with. -/ -def matchTargetStar (mvarId? : Option MVarId) (t : Trie α) : M (Trie α) := do - let (keys, t) ← t.children!.foldr (init := failure) fun (k, c) x => (do - if k == .opaque then - incrementScore 1 - skipEntries c #[k] k.arity - ) <|> x - if let some mvarId := mvarId? then - assignMVar mvarId keys - return t - -/-- Return the possible `Trie α` that come from a `Key.star`, -while keeping track of the `Key.star` assignments. -/ -def matchTreeStars (e : DTExpr) (t : Trie α) : M (Trie α) := do - let {starAssignments, ..} ← get - let mut result := failure - /- The `Key.star` are at the start of the `t.children!`, - so this loops through all of them. -/ - for (k, c) in t.children! do - let .star i := k | break - if let some assignment := starAssignments[i]? then - if e == assignment then - result := (incrementScore e.size *> pure c) <|> result - else - result := (insertStarAssignment i e *> pure c) <|> result - result - -mutual - /-- Return the possible `Trie α` that match with `e`. -/ - partial def matchExpr (e : DTExpr) (t : Trie α) : M (Trie α) := do - if let .star mvarId? := e then - if (← read).unify then - matchTargetStar mvarId? t - else - matchTreeStars e t - else - matchTreeStars e t <|> exactMatch e (findKey t.children!) - - /-- If `e` is not a metavariable, return the possible `Trie α` that exactly match with `e`. -/ - @[specialize] - partial def exactMatch (e : DTExpr) (find? : Key → Option (Trie α)) : M (Trie α) := do - - let findKey (k : Key) (x : Trie α → M (Trie α) := pure) (score := 1) : M (Trie α) := - match find? k with - | none => failure - | some trie => do - incrementScore score - x trie - - let matchArgs (args : Array DTExpr) : Trie α → M (Trie α) := - args.foldlM (fun t e => matchExpr e t) - - match e with - | .opaque => failure - | .const c args => findKey (.const c args.size) (matchArgs args) - | .fvar fvarId args => findKey (.fvar fvarId args.size) (matchArgs args) - | .bvar i args => findKey (.bvar i args.size) (matchArgs args) - | .lit v => findKey (.lit v) - | .sort => findKey .sort - | .lam b => findKey .lam (matchExpr b) 0 - | .forall d b => findKey .forall (matchExpr d >=> matchExpr b) - | .proj n i a args => findKey (.proj n i args.size) (matchExpr a >=> matchArgs args) - | _ => unreachable! - -end - -private partial def getMatchWithScoreAux (d : RefinedDiscrTree α) (e : DTExpr) (unify : Bool) - (allowRootStar : Bool := false) : Array (Array α × Nat) := (do - if e matches .star _ then - guard allowRootStar - d.root.foldl (init := failure) fun x k c => (do - if k == Key.opaque then - GetUnify.incrementScore 1 - let (_, t) ← GetUnify.skipEntries c #[k] k.arity - return t) <|> x - else - GetUnify.exactMatch e d.root.find? - <|> do - guard allowRootStar - let some c := d.root.find? (.star 0) | failure - return c - ).run unify - -end GetUnify - -/-- -Return the results from the `RefinedDiscrTree` that match the given expression, -together with their matching scores, in decreasing order of score. - -Each entry of type `Array α × Nat` corresponds to one pattern. - -If `unify := false`, then metavariables in `e` are treated as opaque variables. -This is for when you don't want to instantiate metavariables in `e`. - -If `allowRootStar := false`, then we don't allow `e` or the matched key in `d` -to be a star pattern. -/ -def getMatchWithScore (d : RefinedDiscrTree α) (e : Expr) (unify : Bool) - (allowRootStar : Bool := false) : MetaM (Array (Array α × Nat)) := do - let e ← mkDTExpr e - let result := GetUnify.getMatchWithScoreAux d e unify allowRootStar - return result.qsort (·.2 > ·.2) - -/-- Similar to `getMatchWithScore`, but also returns matches with prefixes of `e`. -We store the score, followed by the number of ignored arguments. -/ -partial def getMatchWithScoreWithExtra (d : RefinedDiscrTree α) (e : Expr) (unify : Bool) - (allowRootStar : Bool := false) : - MetaM (Array (Array α × Nat × Nat)) := do - let result ← go e 0 - return result.qsort (·.2.1 > ·.2.1) -where - /-- go -/ - go (e : Expr) (numIgnored : Nat) : MetaM (Array (Array α × Nat × Nat)) := do - let result ← getMatchWithScore d e unify allowRootStar - let result := result.map fun (a, b) => (a, b, numIgnored) - match e with - | .app e _ => return (← go e (numIgnored + 1)) ++ result - | _ => return result - - variable {β : Type} {m : Type → Type} [Monad m] /-- Apply a monadic function to the array of values at each node in a `RefinedDiscrTree`. -/ @@ -1147,4 +220,4 @@ def mapArraysM (d : RefinedDiscrTree α) (f : Array α → m (Array β)) : m (Re def mapArrays (d : RefinedDiscrTree α) (f : Array α → Array β) : RefinedDiscrTree β := d.mapArraysM (m := Id) f -end Mathlib.Meta.FunProp.RefinedDiscrTree +end Lean.Meta.RefinedDiscrTree diff --git a/Mathlib/Lean/Meta/RefinedDiscrTree/Basic.lean b/Mathlib/Lean/Meta/RefinedDiscrTree/Basic.lean new file mode 100644 index 0000000000000..b8f2d6390d502 --- /dev/null +++ b/Mathlib/Lean/Meta/RefinedDiscrTree/Basic.lean @@ -0,0 +1,179 @@ +/- +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 + +/-! +# Basic Definitions for `RefinedDiscrTree` +-/ + +namespace Lean.Meta.RefinedDiscrTree + +/-! ## Definitions -/ + +/-- Discrimination tree key. -/ +inductive Key where + /-- A metavariable. This key matches with anything. It stores an index. -/ + | star : Nat → Key + /-- An opaque variable. This key only matches with itself or `Key.star`. -/ + | opaque : Key + /-- A constant. It stores the name and the arity. -/ + | const : Name → Nat → Key + /-- A free variable. It stores the `FVarId` and the arity. -/ + | fvar : FVarId → Nat → Key + /-- A bound variable, from a lambda or forall binder. + It stores the De Bruijn index and the arity. -/ + | bvar : Nat → Nat → Key + /-- A literal. -/ + | lit : Literal → Key + /-- A sort. Universe levels are ignored. -/ + | sort : Key + /-- A lambda function. -/ + | lam : Key + /-- A dependent arrow. -/ + | forall : Key + /-- A projection. It stores the structure name, the projection index and the arity. -/ + | proj : Name → Nat → Nat → Key + deriving Inhabited, BEq, Repr + +private nonrec def Key.hash : Key → UInt64 + | .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 + | .sort => 2411 + | .lam => 4742 + | .«forall» => 9752 + | .proj s i a => mixHash (hash a) <| mixHash (hash s) (hash i) + +instance : Hashable Key := ⟨Key.hash⟩ + +/-- Constructor index used for ordering `Key`. +Note that the index of the star pattern is 0, so that when looking up in a `Trie`, +we can look at the start of the sorted array for all `.star` patterns. -/ +def Key.ctorIdx : Key → Nat + | .star .. => 0 + | .opaque .. => 1 + | .const .. => 2 + | .fvar .. => 3 + | .bvar .. => 4 + | .lit .. => 5 + | .sort => 6 + | .lam => 7 + | .forall => 8 + | .proj .. => 9 + +/-- The order on `Key` used in the `RefinedDiscrTree`. -/ +private def Key.lt : Key → Key → Bool + | .star i₁, .star i₂ => i₁ < i₂ + | .const n₁ a₁, .const n₂ a₂ => Name.quickLt n₁ n₂ || (n₁ == n₂ && a₁ < a₂) + | .fvar f₁ a₁, .fvar f₂ a₂ => Name.quickLt f₁.name f₂.name || (f₁ == f₂ && a₁ < a₂) + | .bvar i₁ a₁, .bvar i₂ a₂ => i₁ < i₂ || (i₁ == i₂ && a₁ < a₂) + | .lit v₁, .lit v₂ => v₁ < v₂ + | .proj s₁ i₁ a₁, .proj s₂ i₂ a₂ => Name.quickLt s₁ s₂ || + (s₁ == s₂ && (i₁ < i₂ || (i₁ == i₂ && a₁ < a₂))) + | k₁, k₂ => k₁.ctorIdx < k₂.ctorIdx + +instance : LT Key := ⟨fun a b => Key.lt a b⟩ +instance (a b : Key) : Decidable (a < b) := inferInstanceAs (Decidable (Key.lt a b)) + +private def Key.format : Key → Format + | .star i => "*" ++ Std.format i + | .opaque => "◾" + | .const k a => "⟨" ++ Std.format k ++ ", " ++ Std.format a ++ "⟩" + | .fvar k a => "⟨" ++ Std.format k.name ++ ", " ++ Std.format a ++ "⟩" + | .lit (Literal.natVal v) => Std.format v + | .lit (Literal.strVal v) => repr v + | .sort => "sort" + | .bvar i a => "⟨" ++ "#" ++ Std.format i ++ ", " ++ Std.format a ++ "⟩" + | .lam => "λ" + | .forall => "∀" + | .proj s i a => "⟨" ++ Std.format s ++"."++ Std.format i ++", "++ Std.format a ++ "⟩" + +instance : ToFormat Key := ⟨Key.format⟩ + +/-- Return the number of arguments that the `Key` takes. -/ +def Key.arity : Key → Nat + | .const _ a => a + | .fvar _ a => a + | .bvar _ a => a + | .lam => 1 + | .forall => 2 + | .proj _ _ a => 1 + a + | _ => 0 + +variable {α : Type} +/-- Discrimination tree trie. See `RefinedDiscrTree`. -/ +inductive Trie (α : Type) where + /-- Map from `Key` to `Trie`. Children is an `Array` of size at least 2, + sorted in increasing order using `Key.lt`. -/ + | node (children : Array (Key × Trie α)) + /-- Sequence of nodes with only one child. `keys` is an `Array` of size at least 1. -/ + | path (keys : Array Key) (child : Trie α) + /-- Leaf of the Trie. `values` is an `Array` of size at least 1. -/ + | values (vs : Array α) +instance : Inhabited (Trie α) := ⟨.node #[]⟩ + +/-- `Trie.path` constructor that only inserts the path if it is non-empty. -/ +def Trie.mkPath (keys : Array Key) (child : Trie α) := + if keys.isEmpty then child else Trie.path keys child + +/-- `Trie` constructor for a single value, taking the keys starting at index `i`. -/ +def Trie.singleton (keys : Array Key) (value : α) (i : Nat) : Trie α := + mkPath keys[i:] (values #[value]) + +/-- `Trie.node` constructor for combining two `Key`, `Trie α` pairs. -/ +def Trie.mkNode2 (k1 : Key) (t1 : Trie α) (k2 : Key) (t2 : Trie α) : Trie α := + if k1 < k2 then + .node #[(k1, t1), (k2, t2)] + else + .node #[(k2, t2), (k1, t1)] + +/-- Return the values from a `Trie α`, assuming that it is a leaf -/ +def Trie.values! : Trie α → Array α + | .values vs => vs + | _ => panic! "expected .values constructor" + +/-- Return the children of a `Trie α`, assuming that it is not a leaf. +The result is sorted by the `Key`'s -/ +def Trie.children! : Trie α → Array (Key × Trie α) + | .node cs => cs + | .path ks c => #[(ks[0]!, mkPath ks[1:] c)] + | .values _ => panic! "did not expect .values constructor" + +private partial def Trie.format [ToFormat α] : Trie α → Format + | .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 + | .path ks c => Format.sbracket (Format.joinSep ks.toList (", ")) + ++ " => " ++ Format.line ++ format c +where + prepend (k : Key) (t : Trie α) : Trie α := match t with + | .path ks c => .path (#[k] ++ ks) c + | t => .path #[k] t +instance [ToFormat α] : ToFormat (Trie α) := ⟨Trie.format⟩ + + +/-- Discrimination tree. It is an index from expressions to values of type `α`. -/ +structure _root_.Lean.Meta.RefinedDiscrTree (α : Type) where + /-- The underlying `PersistentHashMap` of a `RefinedDiscrTree`. -/ + root : PersistentHashMap Key (Trie α) := {} +instance : Inhabited (RefinedDiscrTree α) := ⟨{}⟩ + +private partial def format [ToFormat α] (d : RefinedDiscrTree α) : Format := + let (_, r) := d.root.foldl + (fun (p : Bool × Format) k c => + (false, + p.2 ++ (if p.1 then Format.nil else Format.line) ++ + Format.paren (Std.format k ++ " => " ++ Std.format c))) + (true, Format.nil) + Format.group r + +instance [ToFormat α] : ToFormat (RefinedDiscrTree α) := ⟨format⟩ + +end Lean.Meta.RefinedDiscrTree diff --git a/Mathlib/Lean/Meta/RefinedDiscrTree/Encode.lean b/Mathlib/Lean/Meta/RefinedDiscrTree/Encode.lean new file mode 100644 index 0000000000000..18f7876c84b89 --- /dev/null +++ b/Mathlib/Lean/Meta/RefinedDiscrTree/Encode.lean @@ -0,0 +1,508 @@ +/- +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.Lean.Meta.RefinedDiscrTree.Basic +import Mathlib.Lean.Meta.RefinedDiscrTree.Pi +import Mathlib.Tactic.FunProp.StateList + +/-! +# Encoding an `Expr` as a sequence of `Key`s +-/ + +open Mathlib.Meta.FunProp (StateListT StateListM) + +namespace Lean.Meta.RefinedDiscrTree +variable {α} + +/-- `DTExpr` is a simplified form of `Expr`. +It is the intermediate step for converting from `Expr` to `Array Key`. -/ +inductive DTExpr where + /-- A metavariable. It optionally stores an `MVarId`. -/ + | star : Option MVarId → DTExpr + /-- An opaque variable or a let-expression in the case `WhnfCoreConfig.zeta := false`. -/ + | opaque : DTExpr + /-- A constant. It stores the name and the arguments. -/ + | const : Name → Array DTExpr → DTExpr + /-- A free variable. It stores the `FVarId` and the arguments -/ + | fvar : FVarId → Array DTExpr → DTExpr + /-- A bound variable. It stores the De Bruijn index and the arguments -/ + | bvar : Nat → Array DTExpr → DTExpr + /-- A literal. -/ + | lit : Literal → DTExpr + /-- A sort. -/ + | sort : DTExpr + /-- A lambda function. It stores the body. -/ + | lam : DTExpr → DTExpr + /-- A dependent arrow. It stores the domain and body. -/ + | forall : DTExpr → DTExpr → DTExpr + /-- A projection. It stores the structure name, projection index, struct body and arguments. -/ + | proj : Name → Nat → DTExpr → Array DTExpr → DTExpr +deriving Inhabited, BEq, Repr + +private partial def DTExpr.format : DTExpr → Format + | .star _ => "*" + | .opaque => "◾" + | .const n as => Std.format n ++ formatArgs as + | .fvar n as => Std.format n.name ++ formatArgs as + | .bvar i as => "#" ++ Std.format i ++ formatArgs as + | .lit (Literal.natVal v) => Std.format v + | .lit (Literal.strVal v) => repr v + | .sort => "Sort" + | .lam b => "λ " ++ DTExpr.format b + | .forall d b => DTExpr.format d ++ " → " ++ DTExpr.format b + | .proj _ i a as => DTExpr.format a ++ "." ++ Std.format i ++ formatArgs as +where + formatArgs (as : Array DTExpr) := + if as.isEmpty + then .nil + else " " ++ Format.paren (@Format.joinSep _ ⟨DTExpr.format⟩ as.toList ", ") + +instance : ToFormat DTExpr := ⟨DTExpr.format⟩ + +/-- Return the size of the `DTExpr`. This is used for calculating the matching score when two +expressions are equal. +The score is not incremented at a lambda, which is so that the expressions +`∀ x, p[x]` and `∃ x, p[x]` get the same size. -/ +partial def DTExpr.size : DTExpr → Nat +| .const _ args +| .fvar _ args +| .bvar _ args => args.foldl (init := 1) (· + ·.size) +| .lam b => b.size +| .forall d b => 1 + d.size + b.size +| _ => 1 + +/-- Determine if two `DTExpr`s are equivalent. -/ +def DTExpr.eqv (a b : DTExpr) : Bool := + (go a b).run' {} +where + @[nolint docBlame] + 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₂ + | .fvar n₁ as₁ , .fvar n₂ as₂ => pure (n₁ == n₂) <&&> goArray as₁ as₂ + | .bvar i₁ as₁ , .bvar i₂ as₂ => pure (i₁ == i₂) <&&> goArray as₁ as₂ + | .lit li₁ , .lit li₂ => pure (li₁ == li₂) + | .sort , .sort => pure true + | .lam b₁ , .lam b₂ => go b₁ b₂ + | .forall d₁ b₁ , .forall d₂ b₂ => go d₁ d₂ <&&> go b₁ b₂ + | .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[id₁]? with + | some id => (id == id₂, map) + | none => (true, map.insert id₁ id₂) + | _ , _ => return false + + @[nolint docBlame] + 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 + return false + return true + else + return false + +/-! ## Encoding an Expr -/ + +/-- This state is used to turn the indexing by `MVarId` and `FVarId` in `DTExpr` into +indexing by `Nat` in `Key`. -/ +private structure Flatten.State where + stars : Array MVarId := #[] + +private def getStar (mvarId? : Option MVarId) : StateM Flatten.State Nat := + modifyGet fun s => + match mvarId? with + | some mvarId => match s.stars.findIdx? (· == mvarId) with + | some idx => (idx, s) + | none => (s.stars.size, { s with stars := s.stars.push mvarId }) + | none => (s.stars.size, { s with stars := s.stars.push ⟨.anonymous⟩ }) + +private partial def DTExpr.flattenAux (todo : Array Key) : DTExpr → StateM Flatten.State (Array Key) + | .star i => return todo.push (.star (← getStar i)) + | .opaque => return todo.push .opaque + | .const n as => as.foldlM flattenAux (todo.push (.const n as.size)) + | .fvar f as => as.foldlM flattenAux (todo.push (.fvar f as.size)) + | .bvar i as => as.foldlM flattenAux (todo.push (.bvar i as.size)) + | .lit l => return todo.push (.lit l) + | .sort => return todo.push .sort + | .lam b => flattenAux (todo.push .lam) b + | .«forall» d b => do flattenAux (← flattenAux (todo.push .forall) d) b + | .proj n i e as => do as.foldlM flattenAux (← flattenAux (todo.push (.proj n i as.size)) e) + +/-- Given a `DTExpr`, return the linearized encoding in terms of `Key`, +which is used for `RefinedDiscrTree` indexing. -/ +def DTExpr.flatten (e : DTExpr) (initCapacity := 16) : Array Key := + (DTExpr.flattenAux (.mkEmpty initCapacity) e).run' {} + + + +/-- Return true if `e` is one of the following +- A nat literal (numeral) +- `Nat.zero` +- `Nat.succ x` where `isNumeral x` +- `OfNat.ofNat _ x _` where `isNumeral x` -/ +private partial def isNumeral (e : Expr) : Bool := + if e.isRawNatLit then true + else + let f := e.getAppFn + if !f.isConst then false + else + let fName := f.constName! + if fName == ``Nat.succ && e.getAppNumArgs == 1 then isNumeral e.appArg! + else if fName == ``OfNat.ofNat && e.getAppNumArgs == 3 then isNumeral (e.getArg! 1) + else if fName == ``Nat.zero && e.getAppNumArgs == 0 then true + else false + +/-- Return `some n` if `e` is definitionally equal to the natural number `n`. -/ +private partial def toNatLit? (e : Expr) : Option Literal := + if isNumeral e then + if let some n := loop e then + some (.natVal n) + else + none + else + none +where + loop (e : Expr) : Option Nat := do + let f := e.getAppFn + match f with + | .lit (.natVal n) => return n + | .const fName .. => + if fName == ``Nat.succ && e.getAppNumArgs == 1 then + let r ← loop e.appArg! + return r+1 + else if fName == ``OfNat.ofNat && e.getAppNumArgs == 3 then + loop (e.getArg! 1) + else if fName == ``Nat.zero && e.getAppNumArgs == 0 then + return 0 + else + failure + | _ => failure + +/-- Reduction procedure for the `RefinedDiscrTree` indexing. -/ +partial def reduce (e : Expr) : MetaM Expr := do + let e ← whnfCore e + match (← unfoldDefinition? e) with + | some e => reduce e + | none => match e.etaExpandedStrict? with + | some e => reduce e + | none => return e + +/-- Repeatedly apply reduce while stripping lambda binders and introducing their variables -/ +@[specialize] +partial def lambdaTelescopeReduce {m} [Monad m] [MonadLiftT MetaM m] [MonadControlT MetaM m] + [Nonempty α] (e : Expr) (fvars : List FVarId) + (k : Expr → List FVarId → m α) : m α := do + match ← reduce e with + | .lam n d b bi => + withLocalDecl n bi d fun fvar => + lambdaTelescopeReduce (b.instantiate1 fvar) (fvar.fvarId! :: fvars) k + | e => k e fvars + + + +/-- Check whether the expression is represented by `Key.star`. -/ +def isStar : Expr → Bool + | .mvar .. => true + | .app f _ => isStar f + | _ => false + +/-- Check whether the expression is represented by `Key.star` and has `arg` as an argument. -/ +def isStarWithArg (arg : Expr) : Expr → Bool + | .app f a => if a == arg then isStar f else isStarWithArg arg f + | _ => false + +private partial def DTExpr.hasLooseBVarsAux (i : Nat) : DTExpr → Bool + | .const _ as => as.any (hasLooseBVarsAux i) + | .fvar _ as => as.any (hasLooseBVarsAux i) + | .bvar j as => j ≥ i || as.any (hasLooseBVarsAux i) + | .proj _ _ a as => a.hasLooseBVarsAux i || as.any (hasLooseBVarsAux i) + | .forall d b => d.hasLooseBVarsAux i || b.hasLooseBVarsAux (i+1) + | .lam b => b.hasLooseBVarsAux (i+1) + | _ => false + +/-- Return `true` if `e` contains a loose bound variable. -/ +def DTExpr.hasLooseBVars (e : DTExpr) : Bool := + e.hasLooseBVarsAux 0 + + +namespace MkDTExpr + +private structure Context where + /-- Variables that come from a lambda or forall binder. + The list index gives the De Bruijn index. -/ + bvars : List FVarId := [] + /-- Variables that come from a lambda that has been removed via η-reduction. -/ + forbiddenVars : List FVarId := [] + fvarInContext : FVarId → Bool + +/-- Return for each argument whether it should be ignored. -/ +def getIgnores (fn : Expr) (args : Array Expr) : MetaM (Array Bool) := do + let mut fnType ← inferType fn + let mut result := Array.mkEmpty args.size + let mut j := 0 + for i in [:args.size] do + unless fnType matches .forallE .. do + fnType ← whnfD (fnType.instantiateRevRange j i args) + j := i + let .forallE _ d b bi := fnType | throwError m! "expected function type {indentExpr fnType}" + fnType := b + result := result.push (← isIgnoredArg args[i]! d bi) + return result +where + /-- Return whether the argument should be ignored. -/ + isIgnoredArg (arg domain : Expr) (binderInfo : BinderInfo) : MetaM Bool := do + if domain.isOutParam then + return true + match binderInfo with + | .instImplicit => return true + | .implicit + | .strictImplicit => return !(← isType arg) + | .default => isProof arg + + +@[specialize] +private def withLams {m} [Monad m] [MonadWithReader Context m] + (lambdas : List FVarId) (k : m DTExpr) : m DTExpr := + if lambdas.isEmpty then + k + else do + let e ← withReader (fun c => { c with bvars := lambdas ++ c.bvars }) k + return lambdas.foldl (fun _ => ·.lam) e + + +/-- Return the encoding of `e` as a `DTExpr`. +If `root = false`, then `e` is a strict sub expression of the original expression. -/ +partial def mkDTExprAux (e : Expr) (root : Bool) : ReaderT Context MetaM DTExpr := do + lambdaTelescopeReduce e [] fun e lambdas => + e.withApp fun fn args => do + + let argDTExpr (arg : Expr) (ignore : Bool) : ReaderT Context MetaM DTExpr := + if ignore then pure (.star none) else mkDTExprAux arg false + + let argDTExprs : ReaderT Context MetaM (Array DTExpr) := do + let ignores ← getIgnores fn args + args.mapIdxM fun i arg => + argDTExpr arg ignores[i]! + + match fn with + | .const n _ => + unless root do + if let some (type, lhs, rhs, lambdas') ← reduceHBinOp n args lambdas then + return ← withLams lambdas' do + let type ← mkDTExprAux type false + let lhs ← mkDTExprAux lhs false + let rhs ← mkDTExprAux rhs false + return .const n #[type, type, .star none, .star none, lhs, rhs] + + if let some (type, arg, lambdas') ← reduceUnOp n e.getAppArgs lambdas then + return ← withLams lambdas' do + let type ← mkDTExprAux type false + let arg ← mkDTExprAux arg false + return .const n #[type, .star none, arg] + + /- since `(fun _ => 0) = 0` and `(fun _ => 1) = 1`, + we don't index lambdas before literals -/ + if let some v := toNatLit? e then + return .lit v + withLams lambdas do + return .const n (← argDTExprs) + | .proj s i a => + withLams lambdas do + let a ← argDTExpr a (isClass (← getEnv) s) + return .proj s i a (← argDTExprs) + | .fvar fvarId => + /- we index `fun x => x` as `id` when not at the root -/ + if let fvarId' :: lambdas' := lambdas then + if fvarId' == fvarId && args.isEmpty && !root then + return ← withLams lambdas' do + let type ← mkDTExprAux (← fvarId.getType) false + return .const ``id #[type] + withLams lambdas do + if let some idx := (← read).bvars.findIdx? (· == fvarId) then + return .bvar idx (← argDTExprs) + if (← read).fvarInContext fvarId then + return .fvar fvarId (← argDTExprs) + else + return .opaque + | .mvar mvarId => + /- When the mvarId has arguments, index it with `[*]` instead of `[λ,*]`, + because the star could depend on the bound variables. As a result, + something indexed `[λ,*]` has that the `*` cannot depend on the λ-bound variables -/ + if args.isEmpty then + withLams lambdas do return .star (some mvarId) + else + return .star none + + | .forallE n d b bi => + withLams lambdas do + let d' ← mkDTExprAux d false + let b' ← withLocalDecl n bi d fun fvar => + withReader (fun c => { c with bvars := fvar.fvarId! :: c.bvars }) do + mkDTExprAux (b.instantiate1 fvar) false + return .forall d' b' + | .lit v => withLams lambdas do return .lit v + | .sort _ => withLams lambdas do return .sort + | .letE .. => withLams lambdas do return .opaque + | .lam .. => withLams lambdas do return .opaque + | _ => unreachable! + +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 +a different number of binders, then the resulting De Bruijn indices are offset. +In practice, getting a `.bvar` in a `DTExpr` is very rare, so we exclude such values from the cache. +-/ +instance : MonadCache Expr DTExpr M where + findCached? e := do + let s ← get + return s.find? e + cache e e' := + if e'.hasLooseBVars then + return + else + modify (·.insert e e') + +/-- Return all pairs of body, bound variables that could possibly appear due to η-reduction -/ +@[specialize] +def etaPossibilities (e : Expr) (lambdas : List FVarId) (k : Expr → List FVarId → M α) : M α := + k e lambdas + <|> do + match e, lambdas with + | .app f a, fvarId :: lambdas => + if isStarWithArg (.fvar fvarId) a then + withReader (fun c => { c with forbiddenVars := fvarId :: c.forbiddenVars }) do + etaPossibilities f lambdas k + else + failure + | _, _ => failure + +/-- run `etaPossibilities`, and cache the result if there are multiple possibilities. -/ +@[specialize] +def cacheEtaPossibilities (e original : Expr) (lambdas : List FVarId) + (k : Expr → List FVarId → M DTExpr) : M DTExpr := + match e, lambdas with + | .app _ a, fvarId :: _ => + if isStarWithArg (.fvar fvarId) a then + checkCache original fun _ => + etaPossibilities e lambdas k + else + k e lambdas + | _, _ => k e lambdas + + +/-- Return all encodings of `e` as a `DTExpr`, taking possible η-reductions into account. +If `root = false`, then `e` is a strict sub expression of the original expression. -/ +partial def mkDTExprsAux (original : Expr) (root : Bool) : M DTExpr := do + lambdaTelescopeReduce original [] fun e lambdas => do + + if !root then + if let .const n _ := e.getAppFn then + if let some (type, lhs, rhs, lambdas') ← reduceHBinOp n e.getAppArgs lambdas then + return ← withLams lambdas' do + let type ← mkDTExprsAux type false + let lhs ← mkDTExprsAux lhs false + let rhs ← mkDTExprsAux rhs false + return .const n #[type, type, .star none, .star none, lhs, rhs] + + if let some (type, arg, lambdas') ← reduceUnOp n e.getAppArgs lambdas then + return ← withLams lambdas' do + let type ← mkDTExprsAux type false + let arg ← mkDTExprsAux arg false + return .const n #[type, .star none, arg] + + cacheEtaPossibilities e original lambdas fun e lambdas => + e.withApp fun fn args => do + + let argDTExpr (arg : Expr) (ignore : Bool) : M DTExpr := + if ignore then pure (.star none) else mkDTExprsAux arg false + + let argDTExprs : M (Array DTExpr) := do + let ignores ← getIgnores fn args + args.mapIdxM fun i arg => + argDTExpr arg ignores[i]! + + match fn with + | .const n _ => + unless root do + /- since `(fun _ => 0) = 0` and `(fun _ => 1) = 1`, + we don't index lambdas before nat literals -/ + if let some v := toNatLit? e then + return .lit v + withLams lambdas do + return .const n (← argDTExprs) + | .proj s i a => + withLams lambdas do + let a ← argDTExpr a (isClass (← getEnv) s) + return .proj s i a (← argDTExprs) + | .fvar fvarId => + /- we index `fun x => x` as `id` when not at the root -/ + if let fvarId' :: lambdas' := lambdas then + if fvarId' == fvarId && args.isEmpty && !root then + return ← withLams lambdas' do + let type ← mkDTExprAux (← fvarId.getType) false + return .const ``id #[type] + withLams lambdas do + let c ← read + if let some idx := c.bvars.findIdx? (· == fvarId) then + return .bvar idx (← argDTExprs) + guard !(c.forbiddenVars.contains fvarId) + if c.fvarInContext fvarId then + return .fvar fvarId (← argDTExprs) + else + return .opaque + | .mvar mvarId => + /- When the mvarId has arguments, index it with `[*]` instead of `[λ,*]`, + because the star could depend on the bound variables. As a result, + something indexed `[λ,*]` has that the `*` cannot depend on the λ-bound variables -/ + if args.isEmpty then + withLams lambdas do return .star (some mvarId) + else + return .star none + + | .forallE n d b bi => + withLams lambdas do + let d' ← mkDTExprsAux d false + let b' ← withLocalDecl n bi d fun fvar => + withReader (fun c => { c with bvars := fvar.fvarId! :: c.bvars }) do + mkDTExprsAux (b.instantiate1 fvar) false + return .forall d' b' + | .lit v => withLams lambdas do return .lit v + | .sort _ => withLams lambdas do return .sort + | .letE .. => withLams lambdas do return .opaque + | .lam .. => withLams lambdas do return .opaque + | _ => unreachable! + +end MkDTExpr + +/-- -/ +def DTExpr.isSpecific : DTExpr → Bool + | .star _ + | .const ``Eq #[.star _, .star _, .star _] => false + | _ => true + +/-- Return the encoding of `e` as a `DTExpr`. + +Warning: to account for potential η-reductions of `e`, use `mkDTExprs` instead. + +The argument `fvarInContext` allows you to specify which free variables in `e` will still be +in the context when the `RefinedDiscrTree` is being used for lookup. +It should return true only if the `RefinedDiscrTree` is built and used locally. -/ +def mkDTExpr (e : Expr) + (fvarInContext : FVarId → Bool := fun _ => false) : MetaM DTExpr := + withReducible do (MkDTExpr.mkDTExprAux e true |>.run {fvarInContext}) + +/-- Similar to `mkDTExpr`. +Return all encodings of `e` as a `DTExpr`, taking potential further η-reductions into account. -/ +def mkDTExprs (e : Expr) (onlySpecific : Bool) + (fvarInContext : FVarId → Bool := fun _ => false) : MetaM (List DTExpr) := + withReducible do + let es ← (MkDTExpr.mkDTExprsAux e true).run' {} |>.run {fvarInContext} + return if onlySpecific then es.filter (·.isSpecific) else es + +end Lean.Meta.RefinedDiscrTree diff --git a/Mathlib/Lean/Meta/RefinedDiscrTree/Lookup.lean b/Mathlib/Lean/Meta/RefinedDiscrTree/Lookup.lean new file mode 100644 index 0000000000000..f70e615795728 --- /dev/null +++ b/Mathlib/Lean/Meta/RefinedDiscrTree/Lookup.lean @@ -0,0 +1,187 @@ +/- +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.Lean.Meta.RefinedDiscrTree.Encode + +/-! ## Matching with a RefinedDiscrTree + +We use a very simple unification algorithm. For all star/metavariable patterns in the +`RefinedDiscrTree` and in the target, we store the assignment, and when it is assigned again, +we check that it is the same assignment. +-/ + +open Mathlib.Meta.FunProp (StateListT StateListM) + +namespace Lean.Meta.RefinedDiscrTree +variable {α} + +namespace GetUnify + +/-- If `k` is a key in `children`, return the corresponding `Trie α`. Otherwise return `none`. -/ +def findKey (children : Array (Key × Trie α)) (k : Key) : Option (Trie α) := + (·.2) <$> children.binSearch (k, default) (fun a b => a.1 < b.1) + +private structure Context where + unify : Bool + +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 : Std.HashMap Nat DTExpr := {} + /-- Metavariable assignments for the `Expr.mvar` in the expression. -/ + mvarAssignments : Std.HashMap MVarId (Array Key) := {} + + +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) (x : M (Trie α)) : + Array (Array α × Nat) := + ((x.run { unify }).run {}).toArray.map (fun (t, s) => (t.values!, s.score)) + +/-- Increment the score by `n`. -/ +private def incrementScore (n : Nat) : M Unit := + modify fun s => { s with score := s.score + n } + +/-- Log a metavariable assignment in the `State`. -/ +private def insertStarAssignment (n : Nat) (e : DTExpr) : M Unit := + modify fun s => { s with starAssignments := s.starAssignments.insert n e } + +/-- Log a metavariable assignment in the `State`. -/ +private def assignMVar (mvarId : MVarId) (e : Array Key) : M Unit := do + let { mvarAssignments, .. } ← get + match mvarAssignments[mvarId]? with + | some e' => guard (e == e') + | none => + modify fun s => { s with mvarAssignments := s.mvarAssignments.insert mvarId e } + +/-- Return the possible `Trie α` that match with `n` metavariable. -/ +partial def skipEntries (t : Trie α) (skipped : Array Key) : Nat → M (Array Key × Trie α) + | 0 => pure (skipped, t) + | skip+1 => + t.children!.foldr (init := failure) fun (k, c) x => + (skipEntries c (skipped.push k) (skip + k.arity)) <|> x +/-- Return the possible `Trie α` that match with anything. +We add 1 to the matching score when the key is `.opaque`, +since this pattern is "harder" to match with. -/ +def matchTargetStar (mvarId? : Option MVarId) (t : Trie α) : M (Trie α) := do + let (keys, t) ← t.children!.foldr (init := failure) fun (k, c) x => (do + if k == .opaque then + incrementScore 1 + skipEntries c #[k] k.arity + ) <|> x + if let some mvarId := mvarId? then + assignMVar mvarId keys + return t + +/-- Return the possible `Trie α` that come from a `Key.star`, +while keeping track of the `Key.star` assignments. -/ +def matchTreeStars (e : DTExpr) (t : Trie α) : M (Trie α) := do + let {starAssignments, ..} ← get + let mut result := failure + /- The `Key.star` are at the start of the `t.children!`, + so this loops through all of them. -/ + for (k, c) in t.children! do + let .star i := k | break + if let some assignment := starAssignments[i]? then + if e == assignment then + result := (incrementScore e.size *> pure c) <|> result + else + result := (insertStarAssignment i e *> pure c) <|> result + result + +mutual + /-- Return the possible `Trie α` that match with `e`. -/ + partial def matchExpr (e : DTExpr) (t : Trie α) : M (Trie α) := do + if let .star mvarId? := e then + if (← read).unify then + matchTargetStar mvarId? t + else + matchTreeStars e t + else + matchTreeStars e t <|> exactMatch e (findKey t.children!) + + /-- If `e` is not a metavariable, return the possible `Trie α` that exactly match with `e`. -/ + @[specialize] + partial def exactMatch (e : DTExpr) (find? : Key → Option (Trie α)) : M (Trie α) := do + + let findKey (k : Key) (x : Trie α → M (Trie α) := pure) (score := 1) : M (Trie α) := + match find? k with + | none => failure + | some trie => do + incrementScore score + x trie + + let matchArgs (args : Array DTExpr) : Trie α → M (Trie α) := + args.foldlM (fun t e => matchExpr e t) + + match e with + | .opaque => failure + | .const c args => findKey (.const c args.size) (matchArgs args) + | .fvar fvarId args => findKey (.fvar fvarId args.size) (matchArgs args) + | .bvar i args => findKey (.bvar i args.size) (matchArgs args) + | .lit v => findKey (.lit v) + | .sort => findKey .sort + | .lam b => findKey .lam (matchExpr b) 0 + | .forall d b => findKey .forall (matchExpr d >=> matchExpr b) + | .proj n i a args => findKey (.proj n i args.size) (matchExpr a >=> matchArgs args) + | _ => unreachable! + +end + +private partial def getMatchWithScoreAux (d : RefinedDiscrTree α) (e : DTExpr) (unify : Bool) + (allowRootStar : Bool := false) : Array (Array α × Nat) := (do + if e matches .star _ then + guard allowRootStar + d.root.foldl (init := failure) fun x k c => (do + if k == Key.opaque then + GetUnify.incrementScore 1 + let (_, t) ← GetUnify.skipEntries c #[k] k.arity + return t) <|> x + else + GetUnify.exactMatch e d.root.find? + <|> do + guard allowRootStar + let some c := d.root.find? (.star 0) | failure + return c + ).run unify + +end GetUnify + +/-- +Return the results from the `RefinedDiscrTree` that match the given expression, +together with their matching scores, in decreasing order of score. + +Each entry of type `Array α × Nat` corresponds to one pattern. + +If `unify := false`, then metavariables in `e` are treated as opaque variables. +This is for when you don't want to instantiate metavariables in `e`. + +If `allowRootStar := false`, then we don't allow `e` or the matched key in `d` +to be a star pattern. -/ +def getMatchWithScore (d : RefinedDiscrTree α) (e : Expr) (unify : Bool) + (allowRootStar : Bool := false) : MetaM (Array (Array α × Nat)) := do + let e ← mkDTExpr e + let result := GetUnify.getMatchWithScoreAux d e unify allowRootStar + return result.qsort (·.2 > ·.2) + +/-- Similar to `getMatchWithScore`, but also returns matches with prefixes of `e`. +We store the score, followed by the number of ignored arguments. -/ +partial def getMatchWithScoreWithExtra (d : RefinedDiscrTree α) (e : Expr) (unify : Bool) + (allowRootStar : Bool := false) : + MetaM (Array (Array α × Nat × Nat)) := do + let result ← go e 0 + return result.qsort (·.2.1 > ·.2.1) +where + /-- go -/ + go (e : Expr) (numIgnored : Nat) : MetaM (Array (Array α × Nat × Nat)) := do + let result ← getMatchWithScore d e unify allowRootStar + let result := result.map fun (a, b) => (a, b, numIgnored) + match e with + | .app e _ => return (← go e (numIgnored + 1)) ++ result + | _ => return result + +end Lean.Meta.RefinedDiscrTree diff --git a/Mathlib/Lean/Meta/RefinedDiscrTree/Pi.lean b/Mathlib/Lean/Meta/RefinedDiscrTree/Pi.lean new file mode 100644 index 0000000000000..829a3fd27c71b --- /dev/null +++ b/Mathlib/Lean/Meta/RefinedDiscrTree/Pi.lean @@ -0,0 +1,115 @@ +/- +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 Mathlib.Algebra.Group.Pi.Basic + +/-! +# Reducing Pi instances for indexing in the RefinedDiscrTree +-/ + +namespace Lean.Meta.RefinedDiscrTree + +variable {α} + +/-- Introduce new lambdas by η-expansion. -/ +@[specialize] +def etaExpand (args : Array Expr) (type : Expr) (lambdas : List FVarId) (goalArity : Nat) + (k : Array Expr → List FVarId → MetaM α) : MetaM α := do + if args.size < goalArity then + withLocalDeclD `_η type fun fvar => + etaExpand (args.push fvar) type (fvar.fvarId! :: lambdas) goalArity k + else + k args lambdas + +/-- Normalize an application of a heterogeneous binary operator like `HAdd.hAdd`, using: +- `f = fun x => f x` to increase the arity to 6 +- `(f + g) a = f a + g a` to decrease the arity to 6 +- `(fun x => f x + g x) = f + g` to get rid of any lambdas in front -/ +def reduceHBinOpAux (args : Array Expr) (lambdas : List FVarId) (instH instPi : Name) : + OptionT MetaM (Expr × Expr × Expr × List FVarId) := do + let some (mkApp2 (.const instH' _) type inst) := args[3]? | failure + guard (instH == instH') + if args.size ≤ 6 then + etaExpand args type lambdas 6 fun args lambdas => + distributeLambdas lambdas type args[4]! args[5]! + else + /- use that `(f + g) a = f a + g a` -/ + let mut type := type + let mut inst := inst + let mut lhs := args[4]! + let mut rhs := args[5]! + for arg in args[6:] do + let mkApp3 (.const i _) _ f inst' := inst | return (type, lhs, rhs, lambdas) + unless i == instPi do return (type, lhs, rhs, lambdas) + type := .app f arg + inst := inst' + lhs := .app lhs arg + rhs := .app rhs arg + distributeLambdas lambdas type lhs rhs +where + /-- use that `(fun x => f x + g x) = f + g` -/ + distributeLambdas (lambdas : List FVarId) (type lhs rhs : Expr) : + MetaM (Expr × Expr × Expr × List FVarId) := match lambdas with + | fvarId :: lambdas => do + let decl ← fvarId.getDecl + let type := .forallE decl.userName decl.type (type.abstract #[.fvar fvarId]) decl.binderInfo + let lhs := .lam decl.userName decl.type (lhs.abstract #[.fvar fvarId]) decl.binderInfo + let rhs := .lam decl.userName decl.type (rhs.abstract #[.fvar fvarId]) decl.binderInfo + distributeLambdas lambdas type lhs rhs + | [] => return (type, lhs, rhs, []) + +/-- Normalize an application if the head is `+`, `*`, `-` or `/`. +Optionally return the `(type, lhs, rhs, lambdas)`. -/ +@[inline] def reduceHBinOp (n : Name) (args : Array Expr) (lambdas : List FVarId) : + MetaM (Option (Expr × Expr × Expr × List FVarId)) := + match n with + | ``HAdd.hAdd => reduceHBinOpAux args lambdas ``instHAdd ``Pi.instAdd + | ``HMul.hMul => reduceHBinOpAux args lambdas ``instHMul ``Pi.instMul + | ``HSub.hSub => reduceHBinOpAux args lambdas ``instHSub ``Pi.instSub + | ``HDiv.hDiv => reduceHBinOpAux args lambdas ``instHDiv ``Pi.instDiv + | _ => return none + +/-- Normalize an application of a unary operator like `Inv.inv`, using: +- `f⁻¹ a = (f a)⁻¹` to decrease the arity to 3 +- `(fun x => (f a)⁻¹) = f⁻¹` to get rid of any lambdas in front -/ +def reduceUnOpAux (args : Array Expr) (lambdas : List FVarId) (instPi : Name) : + OptionT MetaM (Expr × Expr × List FVarId) := do + guard (args.size ≥ 3) + let mut type := args[0]! + let mut inst := args[1]! + let mut arg := args[2]! + if args.size == 3 then + distributeLambdas lambdas type arg + else + /- use that `f⁻¹ a = (f a)⁻¹` -/ + for arg' in args[3:] do + let mkApp3 (.const i _) _ f inst' := inst | return (type, arg, lambdas) + unless i == instPi do return (type, arg, lambdas) + type := .app f arg' + inst := inst' + arg := .app arg arg' + distributeLambdas lambdas type arg +where + /-- use that `(fun x => (f x)⁻¹) = f⁻¹` -/ + distributeLambdas (lambdas : List FVarId) (type arg : Expr) : + MetaM (Expr × Expr × List FVarId) := match lambdas with + | fvarId :: lambdas => do + let decl ← fvarId.getDecl + let type := .forallE decl.userName decl.type (type.abstract #[.fvar fvarId]) decl.binderInfo + let arg := .lam decl.userName decl.type (arg.abstract #[.fvar fvarId]) decl.binderInfo + distributeLambdas lambdas type arg + | [] => return (type, arg, []) + +/-- Normalize an application if the head is `⁻¹` or `-`. +Optionally return the `(type, arg, lambdas)`. -/ +@[inline] def reduceUnOp (n : Name) (args : Array Expr) (lambdas : List FVarId) : + MetaM (Option (Expr × Expr × List FVarId)) := + match n with + | ``Neg.neg => reduceUnOpAux args lambdas ``Pi.instNeg + | ``Inv.inv => reduceUnOpAux args lambdas ``Pi.instInv + | _ => return none + +end Lean.Meta.RefinedDiscrTree diff --git a/Mathlib/Tactic/FunProp/Theorems.lean b/Mathlib/Tactic/FunProp/Theorems.lean index abbe43c45aac5..d3089f40eb5b9 100644 --- a/Mathlib/Tactic/FunProp/Theorems.lean +++ b/Mathlib/Tactic/FunProp/Theorems.lean @@ -7,6 +7,7 @@ import Mathlib.Tactic.FunProp.Decl import Mathlib.Tactic.FunProp.Types import Mathlib.Tactic.FunProp.FunctionData import Mathlib.Lean.Meta.RefinedDiscrTree +import Mathlib.Lean.Meta.RefinedDiscrTree.Lookup import Batteries.Data.RBMap.Alter /-! diff --git a/scripts/noshake.json b/scripts/noshake.json index a45672285115c..784455d70e40a 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -357,8 +357,8 @@ "Mathlib.LinearAlgebra.Basic": ["Mathlib.Algebra.BigOperators.Group.Finset"], "Mathlib.LinearAlgebra.AffineSpace.FiniteDimensional": ["Mathlib.Init.Core"], "Mathlib.LinearAlgebra.AffineSpace.Basic": ["Mathlib.Algebra.AddTorsor"], + "Mathlib.Lean.Meta.RefinedDiscrTree.Pi": ["Mathlib.Algebra.Group.Pi.Basic"], "Mathlib.Lean.Meta": ["Batteries.Logic"], - "Mathlib.Lean.Meta.RefinedDiscrTree": ["Mathlib.Algebra.Group.Pi.Basic"], "Mathlib.Lean.Expr.ExtraRecognizers": ["Mathlib.Data.Set.Operations"], "Mathlib.Lean.Expr.Basic": ["Batteries.Logic"], "Mathlib.GroupTheory.MonoidLocalization.Basic": ["Mathlib.Init.Data.Prod"], From 5e5c6da9e047d879ce36de0761405785f64f61fb Mon Sep 17 00:00:00 2001 From: "Tristan F.-R." Date: Thu, 2 Jan 2025 02:10:32 +0000 Subject: [PATCH 058/189] docs(Combinatorics/SimpleGraph/LapMatrix): update theorem name in docs (#20374) Corrects order from `rank_ker_lapMatrix_eq_card_ConnectedComponent` `card_ConnectedComponent_eq_rank_ker_lapMatrix`. Co-authored-by: Tristan F. Co-authored-by: Tristan F.-R. --- Mathlib/Combinatorics/SimpleGraph/LapMatrix.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Combinatorics/SimpleGraph/LapMatrix.lean b/Mathlib/Combinatorics/SimpleGraph/LapMatrix.lean index e609a7dee3f0e..083ee02b783e6 100644 --- a/Mathlib/Combinatorics/SimpleGraph/LapMatrix.lean +++ b/Mathlib/Combinatorics/SimpleGraph/LapMatrix.lean @@ -17,7 +17,7 @@ This module defines the Laplacian matrix of a graph, and proves some of its elem * `SimpleGraph.lapMatrix`: The Laplacian matrix of a simple graph, defined as the difference between the degree matrix and the adjacency matrix. * `isPosSemidef_lapMatrix`: The Laplacian matrix is positive semidefinite. -* `rank_ker_lapMatrix_eq_card_ConnectedComponent`: The number of connected components in `G` is +* `card_ConnectedComponent_eq_rank_ker_lapMatrix`: The number of connected components in `G` is the dimension of the nullspace of its Laplacian matrix. -/ @@ -187,7 +187,7 @@ noncomputable def lapMatrix_ker_basis := end -/-- The number of connected components in `G` is the dimension of the nullspace its Laplacian. -/ +/-- The number of connected components in `G` is the dimension of the nullspace of its Laplacian. -/ theorem card_ConnectedComponent_eq_rank_ker_lapMatrix : Fintype.card G.ConnectedComponent = Module.finrank ℝ (LinearMap.ker (Matrix.toLin' (G.lapMatrix ℝ))) := by classical From ac415859973ceb61aa2f1d24799896641d073121 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 2 Jan 2025 08:45:00 +0000 Subject: [PATCH 059/189] chore: use `ofNat()` around `Real` and `Complex` (#20388) In the places where `no_index` was not present, this makes these lemmas (or mostly, their reverse) work with `simp`. --- Mathlib/Algebra/Star/SelfAdjoint.lean | 4 +--- Mathlib/Analysis/RCLike/Basic.lean | 23 +++++++------------ .../SpecialFunctions/Pow/Complex.lean | 16 +++++-------- Mathlib/Data/Complex/Abs.lean | 3 +-- Mathlib/Data/Complex/Basic.lean | 21 +++++++---------- Mathlib/Data/ENNReal/Basic.lean | 16 ++++--------- Mathlib/Data/NNReal/Defs.lean | 4 +--- Mathlib/Data/Real/ENatENNReal.lean | 4 +--- Mathlib/Data/Real/Irrational.lean | 12 ++++------ 9 files changed, 35 insertions(+), 68 deletions(-) diff --git a/Mathlib/Algebra/Star/SelfAdjoint.lean b/Mathlib/Algebra/Star/SelfAdjoint.lean index db9fa87db3a50..2acc10366d170 100644 --- a/Mathlib/Algebra/Star/SelfAdjoint.lean +++ b/Mathlib/Algebra/Star/SelfAdjoint.lean @@ -199,10 +199,8 @@ variable [Semiring R] [StarRing R] protected theorem natCast (n : ℕ) : IsSelfAdjoint (n : R) := star_natCast _ --- See note [no_index around OfNat.ofNat] @[simp] -protected theorem ofNat (n : ℕ) [n.AtLeastTwo] : - IsSelfAdjoint (no_index (OfNat.ofNat n : R)) := +protected theorem ofNat (n : ℕ) [n.AtLeastTwo] : IsSelfAdjoint (ofNat(n) : R) := .natCast n end Semiring diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index 02a0ac5840750..38d66627fcf9a 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -295,8 +295,7 @@ theorem conj_ofReal (r : ℝ) : conj (r : K) = (r : K) := by theorem conj_nat_cast (n : ℕ) : conj (n : K) = n := map_natCast _ _ --- See note [no_index around OfNat.ofNat] -theorem conj_ofNat (n : ℕ) [n.AtLeastTwo] : conj (no_index (OfNat.ofNat n : K)) = OfNat.ofNat n := +theorem conj_ofNat (n : ℕ) [n.AtLeastTwo] : conj (ofNat(n) : K) = ofNat(n) := map_ofNat _ _ @[rclike_simps, simp] @@ -537,29 +536,23 @@ theorem natCast_re (n : ℕ) : re (n : K) = n := by rw [← ofReal_natCast, ofRe @[simp, rclike_simps, norm_cast] theorem natCast_im (n : ℕ) : im (n : K) = 0 := by rw [← ofReal_natCast, ofReal_im] - --- See note [no_index around OfNat.ofNat] @[simp, rclike_simps] -theorem ofNat_re (n : ℕ) [n.AtLeastTwo] : re (no_index (OfNat.ofNat n) : K) = OfNat.ofNat n := +theorem ofNat_re (n : ℕ) [n.AtLeastTwo] : re (ofNat(n) : K) = ofNat(n) := natCast_re n - --- See note [no_index around OfNat.ofNat] @[simp, rclike_simps] -theorem ofNat_im (n : ℕ) [n.AtLeastTwo] : im (no_index (OfNat.ofNat n) : K) = 0 := +theorem ofNat_im (n : ℕ) [n.AtLeastTwo] : im (ofNat(n) : K) = 0 := natCast_im n --- See note [no_index around OfNat.ofNat] @[rclike_simps, norm_cast] -theorem ofReal_ofNat (n : ℕ) [n.AtLeastTwo] : - ((no_index (OfNat.ofNat n) : ℝ) : K) = OfNat.ofNat n := +theorem ofReal_ofNat (n : ℕ) [n.AtLeastTwo] : ((ofNat(n) : ℝ) : K) = ofNat(n) := ofReal_natCast n theorem ofNat_mul_re (n : ℕ) [n.AtLeastTwo] (z : K) : - re (OfNat.ofNat n * z) = OfNat.ofNat n * re z := by + re (ofNat(n) * z) = ofNat(n) * re z := by rw [← ofReal_ofNat, re_ofReal_mul] theorem ofNat_mul_im (n : ℕ) [n.AtLeastTwo] (z : K) : - im (OfNat.ofNat n * z) = OfNat.ofNat n * im z := by + im (ofNat(n) * z) = ofNat(n) * im z := by rw [← ofReal_ofNat, im_ofReal_mul] @[rclike_simps, norm_cast] @@ -595,11 +588,11 @@ theorem norm_natCast (n : ℕ) : ‖(n : K)‖ = n := by @[simp, rclike_simps, norm_cast] lemma nnnorm_natCast (n : ℕ) : ‖(n : K)‖₊ = n := by simp [nnnorm] @[simp, rclike_simps] -theorem norm_ofNat (n : ℕ) [n.AtLeastTwo] : ‖(no_index (OfNat.ofNat n) : K)‖ = OfNat.ofNat n := +theorem norm_ofNat (n : ℕ) [n.AtLeastTwo] : ‖(ofNat(n) : K)‖ = ofNat(n) := norm_natCast n @[simp, rclike_simps] -lemma nnnorm_ofNat (n : ℕ) [n.AtLeastTwo] : ‖(no_index (OfNat.ofNat n) : K)‖₊ = OfNat.ofNat n := +lemma nnnorm_ofNat (n : ℕ) [n.AtLeastTwo] : ‖(ofNat(n) : K)‖₊ = ofNat(n) := nnnorm_natCast n lemma norm_two : ‖(2 : K)‖ = 2 := norm_ofNat 2 diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean index 49e40ea604b12..1afcd23913d7d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean @@ -105,17 +105,15 @@ lemma cpow_mul_int (x y : ℂ) (n : ℤ) : x ^ (y * n) = (x ^ y) ^ n := by rw [m lemma cpow_nat_mul (x : ℂ) (n : ℕ) (y : ℂ) : x ^ (n * y) = (x ^ y) ^ n := mod_cast cpow_int_mul x n y -/-- See Note [no_index around OfNat.ofNat] -/ lemma cpow_ofNat_mul (x : ℂ) (n : ℕ) [n.AtLeastTwo] (y : ℂ) : - x ^ (no_index (OfNat.ofNat n) * y) = (x ^ y) ^ (OfNat.ofNat n : ℕ) := + x ^ (ofNat(n) * y) = (x ^ y) ^ ofNat(n) := cpow_nat_mul x n y lemma cpow_mul_nat (x y : ℂ) (n : ℕ) : x ^ (y * n) = (x ^ y) ^ n := by rw [mul_comm, cpow_nat_mul] -/-- See Note [no_index around OfNat.ofNat] -/ lemma cpow_mul_ofNat (x y : ℂ) (n : ℕ) [n.AtLeastTwo] : - x ^ (y * no_index (OfNat.ofNat n)) = (x ^ y) ^ (OfNat.ofNat n : ℕ) := + x ^ (y * ofNat(n)) = (x ^ y) ^ ofNat(n) := cpow_mul_nat x y n @[simp, norm_cast] @@ -124,10 +122,9 @@ theorem cpow_natCast (x : ℂ) (n : ℕ) : x ^ (n : ℂ) = x ^ n := by simpa usi @[deprecated (since := "2024-04-17")] alias cpow_nat_cast := cpow_natCast -/-- See Note [no_index around OfNat.ofNat] -/ @[simp] lemma cpow_ofNat (x : ℂ) (n : ℕ) [n.AtLeastTwo] : - x ^ (no_index (OfNat.ofNat n) : ℂ) = x ^ (OfNat.ofNat n : ℕ) := + x ^ (ofNat(n) : ℂ) = x ^ ofNat(n) := cpow_natCast x n theorem cpow_two (x : ℂ) : x ^ (2 : ℂ) = x ^ (2 : ℕ) := cpow_ofNat x 2 @@ -143,10 +140,9 @@ theorem cpow_nat_inv_pow (x : ℂ) {n : ℕ} (hn : n ≠ 0) : (x ^ (n⁻¹ : ℂ rw [← cpow_nat_mul, mul_inv_cancel₀, cpow_one] assumption_mod_cast -/-- See Note [no_index around OfNat.ofNat] -/ @[simp] lemma cpow_ofNat_inv_pow (x : ℂ) (n : ℕ) [n.AtLeastTwo] : - (x ^ ((no_index (OfNat.ofNat n) : ℂ)⁻¹)) ^ (no_index (OfNat.ofNat n) : ℕ) = x := + (x ^ ((ofNat(n) : ℂ)⁻¹)) ^ (ofNat(n) : ℕ) = x := cpow_nat_inv_pow _ (NeZero.ne n) /-- A version of `Complex.cpow_int_mul` with RHS that matches `Complex.cpow_mul`. @@ -168,7 +164,7 @@ lemma cpow_nat_mul' {x : ℂ} {n : ℕ} (hlt : -π < n * x.arg) (hle : n * x.arg lemma cpow_ofNat_mul' {x : ℂ} {n : ℕ} [n.AtLeastTwo] (hlt : -π < OfNat.ofNat n * x.arg) (hle : OfNat.ofNat n * x.arg ≤ π) (y : ℂ) : - x ^ (OfNat.ofNat n * y) = (x ^ (OfNat.ofNat n : ℕ)) ^ y := + x ^ (OfNat.ofNat n * y) = (x ^ ofNat(n)) ^ y := cpow_nat_mul' hlt hle y lemma pow_cpow_nat_inv {x : ℂ} {n : ℕ} (h₀ : n ≠ 0) (hlt : -(π / n) < x.arg) (hle : x.arg ≤ π / n) : @@ -179,7 +175,7 @@ lemma pow_cpow_nat_inv {x : ℂ} {n : ℕ} (h₀ : n ≠ 0) (hlt : -(π / n) < x lemma pow_cpow_ofNat_inv {x : ℂ} {n : ℕ} [n.AtLeastTwo] (hlt : -(π / OfNat.ofNat n) < x.arg) (hle : x.arg ≤ π / OfNat.ofNat n) : - (x ^ (OfNat.ofNat n : ℕ)) ^ ((OfNat.ofNat n : ℂ)⁻¹) = x := + (x ^ ofNat(n)) ^ ((OfNat.ofNat n : ℂ)⁻¹) = x := pow_cpow_nat_inv (NeZero.ne n) hlt hle /-- See also `Complex.pow_cpow_ofNat_inv` for a version that also works for `x * I`, `0 ≤ x`. -/ diff --git a/Mathlib/Data/Complex/Abs.lean b/Mathlib/Data/Complex/Abs.lean index 18f2f48e1b525..4735e0966f985 100644 --- a/Mathlib/Data/Complex/Abs.lean +++ b/Mathlib/Data/Complex/Abs.lean @@ -76,10 +76,9 @@ nonrec theorem abs_of_nonneg {r : ℝ} (h : 0 ≤ r) : Complex.abs r = r := @[simp] theorem abs_natCast (n : ℕ) : Complex.abs n = n := Complex.abs_of_nonneg (Nat.cast_nonneg n) --- See note [no_index around OfNat.ofNat] @[simp] theorem abs_ofNat (n : ℕ) [n.AtLeastTwo] : - Complex.abs (no_index (OfNat.ofNat n : ℂ)) = OfNat.ofNat n := + Complex.abs ofNat(n) = ofNat(n) := abs_natCast n theorem mul_self_abs (z : ℂ) : Complex.abs z * Complex.abs z = normSq z := diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 26342efc5f447..bae81b9aa2c13 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -423,9 +423,7 @@ end noncomputable instance instNNRatCast : NNRatCast ℂ where nnratCast q := ofReal q noncomputable instance instRatCast : RatCast ℂ where ratCast q := ofReal q --- See note [no_index around OfNat.ofNat] -@[simp, norm_cast] lemma ofReal_ofNat (n : ℕ) [n.AtLeastTwo] : - ofReal (no_index (OfNat.ofNat n)) = OfNat.ofNat n := rfl +@[simp, norm_cast] lemma ofReal_ofNat (n : ℕ) [n.AtLeastTwo] : ofReal ofNat(n) = ofNat(n) := rfl @[simp, norm_cast] lemma ofReal_natCast (n : ℕ) : ofReal n = n := rfl @[simp, norm_cast] lemma ofReal_intCast (n : ℤ) : ofReal n = n := rfl @[simp, norm_cast] lemma ofReal_nnratCast (q : ℚ≥0) : ofReal q = q := rfl @@ -434,10 +432,9 @@ noncomputable instance instRatCast : RatCast ℂ where ratCast q := ofReal q @[deprecated (since := "2024-04-17")] alias ofReal_rat_cast := ofReal_ratCast --- See note [no_index around OfNat.ofNat] @[simp] -lemma re_ofNat (n : ℕ) [n.AtLeastTwo] : (no_index (OfNat.ofNat n) : ℂ).re = OfNat.ofNat n := rfl -@[simp] lemma im_ofNat (n : ℕ) [n.AtLeastTwo] : (no_index (OfNat.ofNat n) : ℂ).im = 0 := rfl +lemma re_ofNat (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : ℂ).re = ofNat(n) := rfl +@[simp] lemma im_ofNat (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : ℂ).im = 0 := rfl @[simp, norm_cast] lemma natCast_re (n : ℕ) : (n : ℂ).re = n := rfl @[simp, norm_cast] lemma natCast_im (n : ℕ) : (n : ℂ).im = 0 := rfl @[simp, norm_cast] lemma intCast_re (n : ℤ) : (n : ℂ).re = n := rfl @@ -495,8 +492,7 @@ theorem conj_natCast (n : ℕ) : conj (n : ℂ) = n := map_natCast _ _ @[deprecated (since := "2024-04-17")] alias conj_nat_cast := conj_natCast --- See note [no_index around OfNat.ofNat] -theorem conj_ofNat (n : ℕ) [n.AtLeastTwo] : conj (no_index (OfNat.ofNat n : ℂ)) = OfNat.ofNat n := +theorem conj_ofNat (n : ℕ) [n.AtLeastTwo] : conj (ofNat(n) : ℂ) = ofNat(n) := map_ofNat _ _ -- @[simp] @@ -562,10 +558,9 @@ theorem normSq_ratCast (q : ℚ) : normSq q = q * q := normSq_ofReal _ @[deprecated (since := "2024-04-17")] alias normSq_rat_cast := normSq_ratCast --- See note [no_index around OfNat.ofNat] @[simp] theorem normSq_ofNat (n : ℕ) [n.AtLeastTwo] : - normSq (no_index (OfNat.ofNat n : ℂ)) = OfNat.ofNat n * OfNat.ofNat n := + normSq (ofNat(n) : ℂ) = ofNat(n) * ofNat(n) := normSq_natCast _ @[simp] @@ -796,7 +791,7 @@ lemma div_ratCast (z : ℂ) (x : ℚ) : z / x = ⟨z.re / x, z.im / x⟩ := alias div_rat_cast := div_ratCast lemma div_ofNat (z : ℂ) (n : ℕ) [n.AtLeastTwo] : - z / OfNat.ofNat n = ⟨z.re / OfNat.ofNat n, z.im / OfNat.ofNat n⟩ := + z / ofNat(n) = ⟨z.re / ofNat(n), z.im / ofNat(n)⟩ := div_natCast z n @[simp] lemma div_ofReal_re (z : ℂ) (x : ℝ) : (z / x).re = z.re / x := by rw [div_ofReal] @@ -813,11 +808,11 @@ alias div_rat_cast_im := div_ratCast_im @[simp] lemma div_ofNat_re (z : ℂ) (n : ℕ) [n.AtLeastTwo] : - (z / no_index (OfNat.ofNat n)).re = z.re / OfNat.ofNat n := div_natCast_re z n + (z / ofNat(n)).re = z.re / ofNat(n) := div_natCast_re z n @[simp] lemma div_ofNat_im (z : ℂ) (n : ℕ) [n.AtLeastTwo] : - (z / no_index (OfNat.ofNat n)).im = z.im / OfNat.ofNat n := div_natCast_im z n + (z / ofNat(n)).im = z.im / ofNat(n) := div_natCast_im z n /-! ### Characteristic zero -/ diff --git a/Mathlib/Data/ENNReal/Basic.lean b/Mathlib/Data/ENNReal/Basic.lean index 120674ea5f000..19de739a7c1ee 100644 --- a/Mathlib/Data/ENNReal/Basic.lean +++ b/Mathlib/Data/ENNReal/Basic.lean @@ -359,10 +359,8 @@ lemma coe_ne_one : (r : ℝ≥0∞) ≠ 1 ↔ r ≠ 1 := coe_eq_one.not @[simp, norm_cast] lemma coe_pow (x : ℝ≥0) (n : ℕ) : (↑(x ^ n) : ℝ≥0∞) = x ^ n := rfl --- See note [no_index around OfNat.ofNat] @[simp, norm_cast] -theorem coe_ofNat (n : ℕ) [n.AtLeastTwo] : - ((no_index (OfNat.ofNat n) : ℝ≥0) : ℝ≥0∞) = OfNat.ofNat n := rfl +theorem coe_ofNat (n : ℕ) [n.AtLeastTwo] : ((ofNat(n) : ℝ≥0) : ℝ≥0∞) = ofNat(n) := rfl -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11215): TODO: add lemmas about `OfNat.ofNat` and `<`/`≤` @@ -465,9 +463,7 @@ theorem coe_natCast (n : ℕ) : ((n : ℝ≥0) : ℝ≥0∞) = n := rfl @[simp, norm_cast] lemma ofReal_natCast (n : ℕ) : ENNReal.ofReal n = n := by simp [ENNReal.ofReal] --- See note [no_index around OfNat.ofNat] -@[simp] theorem ofReal_ofNat (n : ℕ) [n.AtLeastTwo] : - ENNReal.ofReal (no_index (OfNat.ofNat n)) = OfNat.ofNat n := +@[simp] theorem ofReal_ofNat (n : ℕ) [n.AtLeastTwo] : ENNReal.ofReal ofNat(n) = ofNat(n) := ofReal_natCast n @[simp, aesop (rule_sets := [finiteness]) safe apply] @@ -476,10 +472,10 @@ theorem natCast_ne_top (n : ℕ) : (n : ℝ≥0∞) ≠ ∞ := WithTop.natCast_n @[simp] theorem natCast_lt_top (n : ℕ) : (n : ℝ≥0∞) < ∞ := WithTop.natCast_lt_top n @[simp, aesop (rule_sets := [finiteness]) safe apply] -lemma ofNat_ne_top {n : ℕ} [Nat.AtLeastTwo n] : no_index (OfNat.ofNat n) ≠ ∞ := natCast_ne_top n +lemma ofNat_ne_top {n : ℕ} [Nat.AtLeastTwo n] : ofNat(n) ≠ ∞ := natCast_ne_top n @[simp] -lemma ofNat_lt_top {n : ℕ} [Nat.AtLeastTwo n] : no_index (OfNat.ofNat n) < ∞ := natCast_lt_top n +lemma ofNat_lt_top {n : ℕ} [Nat.AtLeastTwo n] : ofNat(n) < ∞ := natCast_lt_top n @[simp] theorem top_ne_natCast (n : ℕ) : ∞ ≠ n := WithTop.top_ne_natCast n @@ -493,9 +489,7 @@ theorem toNNReal_nat (n : ℕ) : (n : ℝ≥0∞).toNNReal = n := by theorem toReal_nat (n : ℕ) : (n : ℝ≥0∞).toReal = n := by rw [← ENNReal.ofReal_natCast n, ENNReal.toReal_ofReal (Nat.cast_nonneg _)] --- See note [no_index around OfNat.ofNat] -@[simp] theorem toReal_ofNat (n : ℕ) [n.AtLeastTwo] : - ENNReal.toReal (no_index (OfNat.ofNat n)) = OfNat.ofNat n := +@[simp] theorem toReal_ofNat (n : ℕ) [n.AtLeastTwo] : ENNReal.toReal ofNat(n) = ofNat(n) := toReal_nat n theorem le_coe_iff : a ≤ ↑r ↔ ∃ p : ℝ≥0, a = p ∧ p ≤ r := WithTop.le_coe_iff diff --git a/Mathlib/Data/NNReal/Defs.lean b/Mathlib/Data/NNReal/Defs.lean index c361f32478764..c773102bcabad 100644 --- a/Mathlib/Data/NNReal/Defs.lean +++ b/Mathlib/Data/NNReal/Defs.lean @@ -263,10 +263,8 @@ protected theorem coe_natCast (n : ℕ) : (↑(↑n : ℝ≥0) : ℝ) = n := @[deprecated (since := "2024-04-17")] alias coe_nat_cast := NNReal.coe_natCast --- See note [no_index around OfNat.ofNat] @[simp, norm_cast] -protected theorem coe_ofNat (n : ℕ) [n.AtLeastTwo] : - (no_index (OfNat.ofNat n : ℝ≥0) : ℝ) = OfNat.ofNat n := +protected theorem coe_ofNat (n : ℕ) [n.AtLeastTwo] : ((ofNat(n) : ℝ≥0) : ℝ) = ofNat(n) := rfl @[simp, norm_cast] diff --git a/Mathlib/Data/Real/ENatENNReal.lean b/Mathlib/Data/Real/ENatENNReal.lean index 559f57bd15842..95a23117a7474 100644 --- a/Mathlib/Data/Real/ENatENNReal.lean +++ b/Mathlib/Data/Real/ENatENNReal.lean @@ -49,10 +49,8 @@ theorem toENNReal_top : ((⊤ : ℕ∞) : ℝ≥0∞) = ⊤ := theorem toENNReal_coe (n : ℕ) : ((n : ℕ∞) : ℝ≥0∞) = n := rfl --- See note [no_index around OfNat.ofNat] @[simp, norm_cast] -theorem toENNReal_ofNat (n : ℕ) [n.AtLeastTwo] : - ((no_index (OfNat.ofNat n : ℕ∞)) : ℝ≥0∞) = OfNat.ofNat n := +theorem toENNReal_ofNat (n : ℕ) [n.AtLeastTwo] : ((ofNat(n) : ℕ∞) : ℝ≥0∞) = ofNat(n) := rfl @[simp, norm_cast] diff --git a/Mathlib/Data/Real/Irrational.lean b/Mathlib/Data/Real/Irrational.lean index 12f99f0fcb687..70f269455316c 100644 --- a/Mathlib/Data/Real/Irrational.lean +++ b/Mathlib/Data/Real/Irrational.lean @@ -122,9 +122,8 @@ theorem irrational_sqrt_natCast_iff {n : ℕ} : Irrational (√n) ↔ ¬IsSquare rw [← Rat.isSquare_natCast_iff, ← irrational_sqrt_ratCast_iff_of_nonneg n.cast_nonneg, Rat.cast_natCast] --- See note [no_index around OfNat.ofNat] theorem irrational_sqrt_ofNat_iff {n : ℕ} [n.AtLeastTwo] : - Irrational (√(no_index (OfNat.ofNat n))) ↔ ¬IsSquare (OfNat.ofNat n) := + Irrational √(ofNat(n)) ↔ ¬IsSquare ofNat(n) := irrational_sqrt_natCast_iff theorem Nat.Prime.irrational_sqrt {p : ℕ} (hp : Nat.Prime p) : Irrational (√p) := @@ -147,7 +146,7 @@ unseal Nat.sqrt.iter in example : Irrational √24 := by decide ``` -/ -instance {n : ℕ} [n.AtLeastTwo] : Decidable (Irrational (√(no_index (OfNat.ofNat n)))) := +instance {n : ℕ} [n.AtLeastTwo] : Decidable (Irrational √(ofNat(n))) := decidable_of_iff' _ irrational_sqrt_ofNat_iff instance (n : ℕ) : Decidable (Irrational (√n)) := @@ -188,8 +187,7 @@ theorem ne_zero (h : Irrational x) : x ≠ 0 := mod_cast h.ne_nat 0 theorem ne_one (h : Irrational x) : x ≠ 1 := by simpa only [Nat.cast_one] using h.ne_nat 1 --- See note [no_index around OfNat.ofNat] -@[simp] theorem ne_ofNat (h : Irrational x) (n : ℕ) [n.AtLeastTwo] : x ≠ no_index (OfNat.ofNat n) := +@[simp] theorem ne_ofNat (h : Irrational x) (n : ℕ) [n.AtLeastTwo] : x ≠ ofNat(n) := h.ne_nat n end Irrational @@ -203,9 +201,7 @@ theorem Int.not_irrational (m : ℤ) : ¬Irrational m := fun h => h.ne_int m rfl @[simp] theorem Nat.not_irrational (m : ℕ) : ¬Irrational m := fun h => h.ne_nat m rfl --- See note [no_index around OfNat.ofNat] -@[simp] theorem not_irrational_ofNat (n : ℕ) [n.AtLeastTwo] : - ¬Irrational (no_index (OfNat.ofNat n)) := +@[simp] theorem not_irrational_ofNat (n : ℕ) [n.AtLeastTwo] : ¬Irrational ofNat(n) := n.not_irrational namespace Irrational From e814fc9cda4aed84c5302727cb7d2d0dd45c647e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 2 Jan 2025 09:11:22 +0000 Subject: [PATCH 060/189] feat: dot notation for composition of `AEMeasurable` functions with standard functions (#20348) These are consistent with the statements about composition of `Measurable` functions that appear immediately above. --- .../Function/SpecialFunctions/Basic.lean | 76 ++++++++++++++++++- 1 file changed, 75 insertions(+), 1 deletion(-) diff --git a/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean b/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean index b866f1253eceb..a869548c37016 100644 --- a/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean +++ b/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean @@ -24,7 +24,7 @@ assert_not_exists FiniteDimensional.proper noncomputable section -open NNReal ENNReal +open NNReal ENNReal MeasureTheory namespace Real @@ -165,6 +165,43 @@ theorem Measurable.sqrt : Measurable fun x => √(f x) := end RealComposition +section RealComposition + +open Real + +variable {α : Type*} {m : MeasurableSpace α} {μ : Measure α} {f : α → ℝ} (hf : AEMeasurable f μ) +include hf + +@[measurability, fun_prop] +protected lemma AEMeasurable.exp : AEMeasurable (fun x ↦ exp (f x)) μ := + measurable_exp.comp_aemeasurable hf + +@[measurability, fun_prop] +protected lemma AEMeasurable.log : AEMeasurable (fun x ↦ log (f x)) μ := + measurable_log.comp_aemeasurable hf + +@[measurability, fun_prop] +protected lemma AEMeasurable.cos : AEMeasurable (fun x ↦ cos (f x)) μ := + measurable_cos.comp_aemeasurable hf + +@[measurability, fun_prop] +protected lemma AEMeasurable.sin : AEMeasurable (fun x ↦ sin (f x)) μ := + measurable_sin.comp_aemeasurable hf + +@[measurability, fun_prop] +protected lemma AEMeasurable.cosh : AEMeasurable (fun x ↦ cosh (f x)) μ := + measurable_cosh.comp_aemeasurable hf + +@[measurability, fun_prop] +protected lemma AEMeasurable.sinh : AEMeasurable (fun x ↦ sinh (f x)) μ := + measurable_sinh.comp_aemeasurable hf + +@[measurability, fun_prop] +protected lemma AEMeasurable.sqrt : AEMeasurable (fun x ↦ √(f x)) μ := + continuous_sqrt.measurable.comp_aemeasurable hf + +end RealComposition + section ComplexComposition open Complex @@ -202,6 +239,43 @@ theorem Measurable.clog : Measurable fun x => Complex.log (f x) := end ComplexComposition +section ComplexComposition + +open Complex + +variable {α : Type*} {m : MeasurableSpace α} {μ : Measure α} {f : α → ℂ} (hf : AEMeasurable f μ) +include hf + +@[measurability, fun_prop] +protected lemma AEMeasurable.cexp : AEMeasurable (fun x ↦ exp (f x)) μ := + measurable_exp.comp_aemeasurable hf + +@[measurability, fun_prop] +protected lemma AEMeasurable.ccos : AEMeasurable (fun x ↦ cos (f x)) μ := + measurable_cos.comp_aemeasurable hf + +@[measurability, fun_prop] +protected lemma AEMeasurable.csin : AEMeasurable (fun x ↦ sin (f x)) μ := + measurable_sin.comp_aemeasurable hf + +@[measurability, fun_prop] +protected lemma AEMeasurable.ccosh : AEMeasurable (fun x ↦ cosh (f x)) μ := + measurable_cosh.comp_aemeasurable hf + +@[measurability, fun_prop] +protected lemma AEMeasurable.csinh : AEMeasurable (fun x ↦ sinh (f x)) μ := + measurable_sinh.comp_aemeasurable hf + +@[measurability, fun_prop] +protected lemma AEMeasurable.carg : AEMeasurable (fun x ↦ arg (f x)) μ := + measurable_arg.comp_aemeasurable hf + +@[measurability, fun_prop] +protected lemma AEMeasurable.clog : AEMeasurable (fun x ↦ log (f x)) μ := + measurable_log.comp_aemeasurable hf + +end ComplexComposition + section PowInstances instance Complex.hasMeasurablePow : MeasurablePow ℂ ℂ := From d3cb5ba31b21a3f039e1648207a51a13066a58da Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 2 Jan 2025 09:29:41 +0000 Subject: [PATCH 061/189] chore(AlgebraicClosure): construction with less dependency (#20295) #19853 golfed the construction using a theorem of Isaacs, but introduced more dependency. This PR reduces the dependency again using a cleverer construction that adjoins all roots of every (monic) polynomial, rather than one root of each polynomial. The proof follows https://kconrad.math.uconn.edu/blurbs/galoistheory/algclosureshorter.pdf. --- .../IsAlgClosed/AlgebraicClosure.lean | 146 +++++++++++------- 1 file changed, 88 insertions(+), 58 deletions(-) diff --git a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean index 2c5ad01a7762a..38ba01a9abd26 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ import Mathlib.Algebra.CharP.Algebra -import Mathlib.FieldTheory.Isaacs +import Mathlib.Data.Multiset.Fintype +import Mathlib.FieldTheory.IsAlgClosed.Basic import Mathlib.FieldTheory.SplittingField.Construction /-! @@ -15,11 +16,11 @@ In this file we construct the algebraic closure of a field ## Main Definitions - `AlgebraicClosure k` is an algebraic closure of `k` (in the same universe). - It is constructed by taking the polynomial ring generated by indeterminates `x_f` - corresponding to monic irreducible polynomials `f` with coefficients in `k`, and quotienting - out by a maximal ideal containing every `f(x_f)`. Usually this needs to be repeated countably - many times (see Exercise 1.13 in [atiyah-macdonald]), but using a theorem of Isaacs [Isaacs1980], - once is enough (see https://kconrad.math.uconn.edu/blurbs/galoistheory/algclosure.pdf). + It is constructed by taking the polynomial ring generated by indeterminates + $X_{f,1}, \dots, X_{f,\deg f}$ corresponding to roots of monic irreducible + polynomials `f` with coefficients in `k`, and quotienting out by a maximal + ideal containing every $f - \prod_i (X - X_{f,i})$. The proof follows + https://kconrad.math.uconn.edu/blurbs/galoistheory/algclosureshorter.pdf. ## Tags @@ -36,61 +37,81 @@ variable (k : Type u) [Field k] namespace AlgebraicClosure -open MvPolynomial +/-- The subtype of monic polynomials. -/ +def Monics : Type u := {f : k[X] // f.Monic} -/-- The subtype of monic irreducible polynomials -/ -abbrev MonicIrreducible : Type u := - { f : k[X] // Monic f ∧ Irreducible f } +/-- `Vars k` provides `n` variables $X_{f,1}, \dots, X_{f,n}$ for each monic polynomial +`f : k[X]` of degree `n`. -/ +def Vars : Type u := Σ f : Monics k, Fin f.1.natDegree -/-- Sends a monic irreducible polynomial `f` to `f(x_f)` where `x_f` is a formal indeterminate. -/ -def evalXSelf (f : MonicIrreducible k) : MvPolynomial (MonicIrreducible k) k := - Polynomial.eval₂ MvPolynomial.C (X f) f +variable {k} in +/-- Given a monic polynomial `f : k[X]`, +`subProdXSubC f` is the polynomial $f - \prod_i (X - X_{f,i})$. -/ +def subProdXSubC (f : Monics k) : (MvPolynomial (Vars k) k)[X] := + f.1.map (algebraMap _ _) - ∏ i : Fin f.1.natDegree, (X - C (MvPolynomial.X ⟨f, i⟩)) -/-- The span of `f(x_f)` across monic irreducible polynomials `f` where `x_f` is an -indeterminate. -/ -def spanEval : Ideal (MvPolynomial (MonicIrreducible k) k) := - Ideal.span <| Set.range <| evalXSelf k +/-- The span of all coefficients of `subProdXSubC f` as `f` ranges all polynomials in `k[X]`. -/ +def spanCoeffs : Ideal (MvPolynomial (Vars k) k) := + Ideal.span <| Set.range fun fn : Monics k × ℕ ↦ (subProdXSubC fn.1).coeff fn.2 + +variable {k} + +/-- If a monic polynomial `f : k[X]` splits in `K`, +then it has as many roots (counting multiplicity) as its degree. -/ +def finEquivRoots {K} [Field K] [DecidableEq K] {i : k →+* K} {f : Monics k} (hf : f.1.Splits i) : + Fin f.1.natDegree ≃ (f.1.map i).roots.toEnumFinset := + .symm <| Finset.equivFinOfCardEq <| by + rwa [← splits_id_iff_splits, splits_iff_card_roots, + ← Multiset.card_toEnumFinset, f.2.natDegree_map] at hf + +lemma Monics.splits_finsetProd {s : Finset (Monics k)} {f : Monics k} (hf : f ∈ s) : + f.1.Splits (algebraMap k (SplittingField (∏ f ∈ s, f.1))) := + (splits_prod_iff _ fun j _ ↦ j.2.ne_zero).1 (SplittingField.splits _) _ hf 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`. -/ -def toSplittingField (s : Finset (MonicIrreducible k)) : - MvPolynomial (MonicIrreducible k) k →ₐ[k] SplittingField (∏ x ∈ s, x : k[X]) := - MvPolynomial.aeval fun f => - if hf : f ∈ s then - rootOfSplits _ - ((splits_prod_iff _ fun (j : MonicIrreducible k) _ => j.2.2.ne_zero).1 - (SplittingField.splits _) f hf) - (mt isUnit_iff_degree_eq_zero.2 f.2.2.not_unit) - else 37 - -theorem toSplittingField_evalXSelf {s : Finset (MonicIrreducible k)} {f} (hf : f ∈ s) : - toSplittingField k s (evalXSelf k f) = 0 := by - rw [toSplittingField, evalXSelf, ← AlgHom.coe_toRingHom, hom_eval₂, AlgHom.coe_toRingHom, - MvPolynomial.aeval_X, dif_pos hf, ← MvPolynomial.algebraMap_eq, AlgHom.comp_algebraMap] - exact map_rootOfSplits _ _ _ - -theorem spanEval_ne_top : spanEval k ≠ ⊤ := by - rw [Ideal.ne_top_iff_one, spanEval, Ideal.span, ← Set.image_univ, +/-- Given a finite set of monic polynomials, construct an algebra homomorphism +to the splitting field of the product of the polynomials +sending indeterminates $X_{f_i}$ to the distinct roots of `f`. -/ +def toSplittingField (s : Finset (Monics k)) : + MvPolynomial (Vars k) k →ₐ[k] SplittingField (∏ f ∈ s, f.1) := + MvPolynomial.aeval fun fi ↦ + if hf : fi.1 ∈ s then (finEquivRoots (Monics.splits_finsetProd hf) fi.2).1.1 else 37 + +theorem toSplittingField_coeff {s : Finset (Monics k)} {f} (h : f ∈ s) (n) : + toSplittingField s ((subProdXSubC f).coeff n) = 0 := by + classical + simp_rw [← AlgHom.coe_toRingHom, ← coeff_map, subProdXSubC, Polynomial.map_sub, + Polynomial.map_prod, Polynomial.map_sub, map_X, map_C, toSplittingField, + AlgHom.coe_toRingHom, MvPolynomial.aeval_X, dif_pos h, + ← (finEquivRoots (Monics.splits_finsetProd h)).symm.prod_comp, Equiv.apply_symm_apply] + rw [Finset.prod_coe_sort (f := fun x : _ × ℕ ↦ X - C x.1), (Multiset.toEnumFinset _) + |>.prod_eq_multiset_prod, ← Function.comp_def (X - C ·) Prod.fst, ← Multiset.map_map, + Multiset.map_toEnumFinset_fst, map_map, AlgHom.comp_algebraMap] + conv in map _ _ => rw [eq_prod_roots_of_splits (Monics.splits_finsetProd h)] + rw [f.2, map_one, C_1, one_mul, sub_self, coeff_zero] + +variable (k) + +theorem spanCoeffs_ne_top : spanCoeffs k ≠ ⊤ := by + rw [Ideal.ne_top_iff_one, spanCoeffs, Ideal.span, ← Set.image_univ, Finsupp.mem_span_image_iff_linearCombination] rintro ⟨v, _, hv⟩ - replace hv := congr_arg (toSplittingField k v.support) hv + classical + replace hv := congr_arg (toSplittingField <| v.support.image Prod.fst) hv rw [map_one, Finsupp.linearCombination_apply, Finsupp.sum, map_sum, Finset.sum_eq_zero] at hv · exact zero_ne_one hv intro j hj - rw [smul_eq_mul, map_mul, toSplittingField_evalXSelf _ (s := v.support) hj, - mul_zero] + rw [smul_eq_mul, map_mul, toSplittingField_coeff (Finset.mem_image_of_mem _ hj), mul_zero] /-- A random maximal ideal that contains `spanEval k` -/ -def maxIdeal : Ideal (MvPolynomial (MonicIrreducible k) k) := - Classical.choose <| Ideal.exists_le_maximal _ <| spanEval_ne_top k +def maxIdeal : Ideal (MvPolynomial (Vars k) k) := + Classical.choose <| Ideal.exists_le_maximal _ <| spanCoeffs_ne_top k instance maxIdeal.isMaximal : (maxIdeal k).IsMaximal := - (Classical.choose_spec <| Ideal.exists_le_maximal _ <| spanEval_ne_top k).1 + (Classical.choose_spec <| Ideal.exists_le_maximal _ <| spanCoeffs_ne_top k).1 -theorem le_maxIdeal : spanEval k ≤ maxIdeal k := - (Classical.choose_spec <| Ideal.exists_le_maximal _ <| spanEval_ne_top k).2 +theorem le_maxIdeal : spanCoeffs k ≤ maxIdeal k := + (Classical.choose_spec <| Ideal.exists_le_maximal _ <| spanCoeffs_ne_top k).2 end AlgebraicClosure @@ -99,7 +120,7 @@ open AlgebraicClosure in each polynomial over the field. -/ @[stacks 09GT] def AlgebraicClosure : Type u := - MvPolynomial (MonicIrreducible k) k ⧸ maxIdeal k + MvPolynomial (Vars k) k ⧸ maxIdeal k namespace AlgebraicClosure @@ -134,6 +155,15 @@ instance instField : Field (AlgebraicClosure k) where qsmul_def q x := Quotient.inductionOn x fun p ↦ congr_arg Quotient.mk'' <| by ext; simp [MvPolynomial.algebraMap_eq, Rat.smul_def] +theorem Monics.map_eq_prod {f : Monics k} : + f.1.map (algebraMap k (AlgebraicClosure k)) = + ∏ i, map (Ideal.Quotient.mk <| maxIdeal k) (X - C (MvPolynomial.X ⟨f, i⟩)) := by + ext + -- erw due to `AlgebraicClosure k` not being reducibly defeq to the quotient by `maxIdeal k` + erw [← Ideal.Quotient.mk_comp_algebraMap, ← map_map, ← Polynomial.map_prod, ← sub_eq_zero, + ← coeff_sub, ← Polynomial.map_sub, ← subProdXSubC, coeff_map, Ideal.Quotient.eq_zero_iff_mem] + refine le_maxIdeal _ (Ideal.subset_span ⟨⟨f, _⟩, rfl⟩) + instance isAlgebraic : Algebra.IsAlgebraic k (AlgebraicClosure k) := ⟨fun z => IsIntegral.isAlgebraic <| by @@ -142,17 +172,17 @@ instance isAlgebraic : Algebra.IsAlgebraic k (AlgebraicClosure k) := induction p using MvPolynomial.induction_on generalizing z with | h_C => exact isIntegral_algebraMap | h_add _ _ ha hb => exact (ha _ rfl).add (hb _ rfl) - | h_X p f ih => - refine @IsIntegral.mul k _ _ _ _ _ (Ideal.Quotient.mk (maxIdeal k) _) (ih _ rfl) ?_ - refine ⟨f, f.2.1, ?_⟩ - erw [algebraMap, ← hom_eval₂, Ideal.Quotient.eq_zero_iff_mem] - exact le_maxIdeal k (Ideal.subset_span ⟨f, rfl⟩)⟩ - -instance : IsAlgClosure k (AlgebraicClosure k) := .of_exist_roots fun f hfm hfi ↦ - ⟨Ideal.Quotient.mk _ <| MvPolynomial.X (⟨f, hfm, hfi⟩ : MonicIrreducible k), by - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [aeval_def, ← hom_eval₂, Ideal.Quotient.eq_zero_iff_mem] - exact le_maxIdeal k (Ideal.subset_span <| ⟨_, rfl⟩)⟩ + | h_X p fi ih => + rw [map_mul] + refine (ih _ rfl).mul ⟨_, fi.1.2, ?_⟩ + simp_rw [← eval_map, Monics.map_eq_prod, eval_prod, Polynomial.map_sub, eval_sub] + apply Finset.prod_eq_zero (Finset.mem_univ fi.2) + erw [map_C, eval_C] + simp⟩ + +instance : IsAlgClosure k (AlgebraicClosure k) := .of_splits fun f hf _ ↦ by + rw [show f = (⟨f, hf⟩ : Monics k) from rfl, ← splits_id_iff_splits, Monics.map_eq_prod] + exact splits_prod _ fun _ _ ↦ (splits_id_iff_splits _).mpr (splits_X_sub_C _) instance isAlgClosed : IsAlgClosed (AlgebraicClosure k) := IsAlgClosure.isAlgClosed k From 680465ca1f778273b8f84e44162f80afba4cfabe Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 2 Jan 2025 09:59:42 +0000 Subject: [PATCH 062/189] feat: an `RCLike` field is a real-inner product space of dim at most two (#20321) Co-authored-by: Heather Macbeth --- Mathlib/Analysis/InnerProductSpace/Basic.lean | 34 +++++++++++++++---- Mathlib/Analysis/RCLike/Lemmas.lean | 29 +++++++++++++--- 2 files changed, 52 insertions(+), 11 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index cd34649342a00..3bc54c60335ff 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -860,14 +860,20 @@ variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y /-- A general inner product implies a real inner product. This is not registered as an instance -since it creates problems with the case `𝕜 = ℝ`. -/ +since `𝕜` does not appear in the return type `Inner ℝ E`. -/ def Inner.rclikeToReal : Inner ℝ E where inner x y := re ⟪x, y⟫ -/-- A general inner product space structure implies a real inner product structure. This is not -registered as an instance since it creates problems with the case `𝕜 = ℝ`, but in can be used in a -proof to obtain a real inner product space structure from a given `𝕜`-inner product space -structure. -/ -def InnerProductSpace.rclikeToReal : InnerProductSpace ℝ E := +/-- A general inner product space structure implies a real inner product structure. + +This is not registered as an instance since +* `𝕜` does not appear in the return type `InnerProductSpace ℝ E`, +* It is likely to create instance diamonds, as it builds upon the diamond-prone + `NormedSpace.restrictScalars`. + +However, it can be used in a proof to obtain a real inner product space structure from a given +`𝕜`-inner product space structure. -/ +-- See note [reducible non instances] +abbrev InnerProductSpace.rclikeToReal : InnerProductSpace ℝ E := { Inner.rclikeToReal 𝕜 E, NormedSpace.restrictScalars ℝ 𝕜 E with @@ -905,6 +911,22 @@ protected theorem Complex.inner (w z : ℂ) : ⟪w, z⟫_ℝ = (conj w * z).re : end RCLikeToReal +/-- An `RCLike` field is a real inner product space. -/ +noncomputable instance RCLike.toInnerProductSpaceReal : InnerProductSpace ℝ 𝕜 where + __ := Inner.rclikeToReal 𝕜 𝕜 + norm_sq_eq_inner := norm_sq_eq_inner + conj_symm x y := inner_re_symm .. + add_left x y z := + show re (_ * _) = re (_ * _) + re (_ * _) by simp only [map_add, mul_re, conj_re, conj_im]; ring + smul_left x y r := + show re (_ * _) = _ * re (_ * _) by + simp only [mul_re, conj_re, conj_im, conj_trivial, smul_re, smul_im]; ring + +-- The instance above does not create diamonds for concrete `𝕜`: +example : (innerProductSpace : InnerProductSpace ℝ ℝ) = RCLike.toInnerProductSpaceReal := rfl +example : + (instInnerProductSpaceRealComplex : InnerProductSpace ℝ ℂ) = RCLike.toInnerProductSpaceReal := rfl + section Continuous variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] diff --git a/Mathlib/Analysis/RCLike/Lemmas.lean b/Mathlib/Analysis/RCLike/Lemmas.lean index 2fa56df7b41a0..1990ad092ed9a 100644 --- a/Mathlib/Analysis/RCLike/Lemmas.lean +++ b/Mathlib/Analysis/RCLike/Lemmas.lean @@ -8,6 +8,7 @@ import Mathlib.Analysis.RCLike.Basic /-! # Further lemmas about `RCLike` -/ +open scoped Finset variable {K E : Type*} [RCLike K] @@ -18,6 +19,28 @@ theorem ofReal_eval (p : ℝ[X]) (x : ℝ) : (↑(p.eval x) : K) = aeval (↑x) end Polynomial +variable (K) in +lemma RCLike.span_one_I : Submodule.span ℝ (M := K) {1, I} = ⊤ := by + suffices ∀ x : K, ∃ a b : ℝ, a • 1 + b • I = x by + simpa [Submodule.eq_top_iff', Submodule.mem_span_pair] + exact fun x ↦ ⟨re x, im x, by simp [real_smul_eq_coe_mul]⟩ + +variable (K) in +lemma RCLike.rank_le_two : Module.rank ℝ K ≤ 2 := + calc + _ = Module.rank ℝ ↥(Submodule.span ℝ ({1, I} : Set K)) := by rw [span_one_I]; simp + _ ≤ #({1, I} : Finset K) := by + -- TODO: `simp` doesn't rewrite inside the type argument to `Module.rank`, but `rw` does. + -- We should introduce `Submodule.rank` to fix this. + have := rank_span_finset_le (R := ℝ) (M := K) {1, I} + rw [Finset.coe_pair] at this + simpa [span_one_I] using this + _ ≤ 2 := mod_cast Finset.card_le_two + +variable (K) in +lemma RCLike.finrank_le_two : Module.finrank ℝ K ≤ 2 := + Module.finrank_le_of_rank_le <| rank_le_two _ + namespace FiniteDimensional open RCLike @@ -27,11 +50,7 @@ This instance generates a type-class problem with a metavariable `?m` that shoul `RCLike ?m`. Since this can only be satisfied by `ℝ` or `ℂ`, this does not cause problems. -/ /-- An `RCLike` field is finite-dimensional over `ℝ`, since it is spanned by `{1, I}`. -/ -instance rclike_to_real : FiniteDimensional ℝ K := - ⟨{1, I}, by - suffices ∀ x : K, ∃ a b : ℝ, a • 1 + b • I = x by - simpa [Submodule.eq_top_iff', Submodule.mem_span_pair] - exact fun x ↦ ⟨re x, im x, by simp [real_smul_eq_coe_mul]⟩⟩ +instance rclike_to_real : FiniteDimensional ℝ K := ⟨{1, I}, by simp [span_one_I]⟩ variable (K E) variable [NormedAddCommGroup E] [NormedSpace K E] From 38800e147b0bd6d9173e862c744609a4d8c18d1d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 2 Jan 2025 10:14:07 +0000 Subject: [PATCH 063/189] chore: protect `Measurable.exp` (#20390) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit … and adjacent theorems. This avoids `exp` being ambiguous between `Real.exp` and `Measurable.exp`. --- .../Function/SpecialFunctions/Basic.lean | 33 ++++++++----------- 1 file changed, 14 insertions(+), 19 deletions(-) diff --git a/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean b/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean index a869548c37016..e072184cd1af6 100644 --- a/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean +++ b/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean @@ -136,32 +136,27 @@ variable {α : Type*} {m : MeasurableSpace α} {f : α → ℝ} (hf : Measurable include hf @[measurability] -theorem Measurable.exp : Measurable fun x => Real.exp (f x) := +protected theorem Measurable.exp : Measurable fun x => Real.exp (f x) := Real.measurable_exp.comp hf @[measurability] -theorem Measurable.log : Measurable fun x => log (f x) := +protected theorem Measurable.log : Measurable fun x => log (f x) := measurable_log.comp hf @[measurability] -theorem Measurable.cos : Measurable fun x => Real.cos (f x) := - Real.measurable_cos.comp hf +protected theorem Measurable.cos : Measurable fun x ↦ cos (f x) := measurable_cos.comp hf @[measurability] -theorem Measurable.sin : Measurable fun x => Real.sin (f x) := - Real.measurable_sin.comp hf +protected theorem Measurable.sin : Measurable fun x ↦ sin (f x) := measurable_sin.comp hf @[measurability] -theorem Measurable.cosh : Measurable fun x => Real.cosh (f x) := - Real.measurable_cosh.comp hf +protected theorem Measurable.cosh : Measurable fun x ↦ cosh (f x) := measurable_cosh.comp hf @[measurability] -theorem Measurable.sinh : Measurable fun x => Real.sinh (f x) := - Real.measurable_sinh.comp hf +protected theorem Measurable.sinh : Measurable fun x ↦ sinh (f x) := measurable_sinh.comp hf @[measurability] -theorem Measurable.sqrt : Measurable fun x => √(f x) := - continuous_sqrt.measurable.comp hf +protected theorem Measurable.sqrt : Measurable fun x => √(f x) := continuous_sqrt.measurable.comp hf end RealComposition @@ -210,31 +205,31 @@ variable {α : Type*} {m : MeasurableSpace α} {f : α → ℂ} (hf : Measurable include hf @[measurability] -theorem Measurable.cexp : Measurable fun x => Complex.exp (f x) := +protected theorem Measurable.cexp : Measurable fun x => Complex.exp (f x) := Complex.measurable_exp.comp hf @[measurability] -theorem Measurable.ccos : Measurable fun x => Complex.cos (f x) := +protected theorem Measurable.ccos : Measurable fun x => Complex.cos (f x) := Complex.measurable_cos.comp hf @[measurability] -theorem Measurable.csin : Measurable fun x => Complex.sin (f x) := +protected theorem Measurable.csin : Measurable fun x => Complex.sin (f x) := Complex.measurable_sin.comp hf @[measurability] -theorem Measurable.ccosh : Measurable fun x => Complex.cosh (f x) := +protected theorem Measurable.ccosh : Measurable fun x => Complex.cosh (f x) := Complex.measurable_cosh.comp hf @[measurability] -theorem Measurable.csinh : Measurable fun x => Complex.sinh (f x) := +protected theorem Measurable.csinh : Measurable fun x => Complex.sinh (f x) := Complex.measurable_sinh.comp hf @[measurability] -theorem Measurable.carg : Measurable fun x => arg (f x) := +protected theorem Measurable.carg : Measurable fun x => arg (f x) := measurable_arg.comp hf @[measurability] -theorem Measurable.clog : Measurable fun x => Complex.log (f x) := +protected theorem Measurable.clog : Measurable fun x => Complex.log (f x) := measurable_log.comp hf end ComplexComposition From 34a6e52b33c6682d11b29adae0a2b98638c498c9 Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Thu, 2 Jan 2025 10:26:20 +0000 Subject: [PATCH 064/189] feat(Data/BitVec): Equivalences between `Fin` and `ZMod` and `BitVec` (#20322) --- Mathlib/Data/BitVec.lean | 11 +++++++++++ Mathlib/Data/ZMod/Basic.lean | 5 +++++ 2 files changed, 16 insertions(+) diff --git a/Mathlib/Data/BitVec.lean b/Mathlib/Data/BitVec.lean index 332dd1a74ab95..1c66535fd68f2 100644 --- a/Mathlib/Data/BitVec.lean +++ b/Mathlib/Data/BitVec.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Harun Khan, Alex Keizer -/ import Mathlib.Algebra.Ring.InjSurj +import Mathlib.Algebra.Ring.Equiv import Mathlib.Data.ZMod.Defs /-! @@ -89,4 +90,14 @@ instance : CommRing (BitVec w) := toFin_zero toFin_one toFin_add toFin_mul toFin_neg toFin_sub toFin_nsmul toFin_zsmul toFin_pow toFin_natCast toFin_intCast +/-- The ring `BitVec m` is isomorphic to `Fin (2 ^ m)`. -/ +@[simps] +def equivFin {m : ℕ} : BitVec m ≃+* Fin (2 ^ m) where + toFun a := a.toFin + invFun a := ofFin a + left_inv _ := rfl + right_inv _ := rfl + map_mul' _ _ := rfl + map_add' _ _ := rfl + end BitVec diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 05692bc1135dc..33ba9dfef1f77 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -39,6 +39,11 @@ open Function ZMod namespace ZMod +/-- For non-zero `n : ℕ`, the ring `Fin n` is equivalent to `ZMod n`. -/ +def finEquiv : ∀ (n : ℕ) [NeZero n], Fin n ≃+* ZMod n + | 0, h => (h.ne _ rfl).elim + | _ + 1, _ => .refl _ + instance charZero : CharZero (ZMod 0) := inferInstanceAs (CharZero ℤ) /-- `val a` is a natural number defined as: From c81c0706e20e51611bc65fcbfeb6bc744dd3ba46 Mon Sep 17 00:00:00 2001 From: Patrick Massot Date: Thu, 2 Jan 2025 10:26:22 +0000 Subject: [PATCH 065/189] chore: remove mention of non-existent definition (#20391) Fixes #10059 --- Mathlib/Data/Finset/Defs.lean | 3 --- 1 file changed, 3 deletions(-) diff --git a/Mathlib/Data/Finset/Defs.lean b/Mathlib/Data/Finset/Defs.lean index 7acf539ddfefc..ff2e09b64a5eb 100644 --- a/Mathlib/Data/Finset/Defs.lean +++ b/Mathlib/Data/Finset/Defs.lean @@ -30,9 +30,6 @@ More information can be found in `Mathlib/Algebra/BigOperators/Group/Finset.lean Finsets are directly used to define fintypes in Lean. A `Fintype α` instance for a type `α` consists of a universal `Finset α` containing every term of `α`, called `univ`. See `Mathlib/Data/Fintype/Basic.lean`. -There is also `univ'`, the noncomputable partner to `univ`, -which is defined to be `α` as a finset if `α` is finite, -and the empty finset otherwise. See `Mathlib/Data/Fintype/Basic.lean`. `Finset.card`, the size of a finset is defined in `Mathlib/Data/Finset/Card.lean`. This is then used to define `Fintype.card`, the size of a type. From 6f76ae273407e5fd4054f3d8307fe842a2b85d05 Mon Sep 17 00:00:00 2001 From: Nailin Guan <150537269+Thmoas-Guan@users.noreply.github.com> Date: Thu, 2 Jan 2025 11:03:37 +0000 Subject: [PATCH 066/189] feat(Topology/Group): continuous isomorphism (#16991) Define the continuous isomorphisms of multiplicative and additive topological group. --- .../Topology/Algebra/ContinuousMonoidHom.lean | 279 +++++++++++++++++- 1 file changed, 278 insertions(+), 1 deletion(-) diff --git a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean index f7373a1da6da0..8c4ddb97a123d 100644 --- a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean +++ b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2022 Thomas Browning. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Thomas Browning +Authors: Thomas Browning, Nailin Guan -/ import Mathlib.Topology.Algebra.Equicontinuity import Mathlib.Topology.Algebra.Group.Compact @@ -20,6 +20,8 @@ This file defines the space of continuous homomorphisms between two topological * `ContinuousAddMonoidHom A B`: The continuous additive homomorphisms `A →+ B`. -/ +section + open Function Topology open scoped Pointwise @@ -413,3 +415,278 @@ theorem locallyCompactSpace_of_hasBasis (V : ℕ → Set Y) end LocallyCompact end ContinuousMonoidHom + +end + +section + +/-! + +# Continuous MulEquiv + +This section defines the space of continuous isomorphisms between two topological groups. + +## Main definitions + +-/ + +universe u v + +variable (G : Type u) [TopologicalSpace G] (H : Type v) [TopologicalSpace H] + +/-- The structure of two-sided continuous isomorphisms between additive groups. +Note that both the map and its inverse have to be continuous. -/ +structure ContinuousAddEquiv [Add G] [Add H] extends G ≃+ H , G ≃ₜ H + +/-- The structure of two-sided continuous isomorphisms between groups. +Note that both the map and its inverse have to be continuous. -/ +@[to_additive "The structure of two-sided continuous isomorphisms between additive groups. +Note that both the map and its inverse have to be continuous."] +structure ContinuousMulEquiv [Mul G] [Mul H] extends G ≃* H , G ≃ₜ H + +/-- The homeomorphism induced from a two-sided continuous isomorphism of groups. -/ +add_decl_doc ContinuousMulEquiv.toHomeomorph + +/-- The homeomorphism induced from a two-sided continuous isomorphism additive groups. -/ +add_decl_doc ContinuousAddEquiv.toHomeomorph + +@[inherit_doc] +infixl:25 " ≃ₜ* " => ContinuousMulEquiv + +@[inherit_doc] +infixl:25 " ≃ₜ+ " => ContinuousAddEquiv + +section + +namespace ContinuousMulEquiv + +variable {M N : Type*} [TopologicalSpace M] [TopologicalSpace N] [Mul M] [Mul N] + +section coe + +@[to_additive] +instance : EquivLike (M ≃ₜ* N) M N where + coe f := f.toFun + inv f := f.invFun + left_inv f := f.left_inv + right_inv f := f.right_inv + coe_injective' f g h₁ h₂ := by + cases f + cases g + congr + exact MulEquiv.ext_iff.mpr (congrFun h₁) + +@[to_additive] +instance : MulEquivClass (M ≃ₜ* N) M N where + map_mul f := f.map_mul' + +@[to_additive] +instance : HomeomorphClass (M ≃ₜ* N) M N where + map_continuous f := f.continuous_toFun + inv_continuous f := f.continuous_invFun + +/-- Two continuous multiplicative isomorphisms agree if they are defined by the +same underlying function. -/ +@[to_additive (attr := ext) + "Two continuous additive isomorphisms agree if they are defined by the same underlying function."] +theorem ext {f g : M ≃ₜ* N} (h : ∀ x, f x = g x) : f = g := + DFunLike.ext f g h + +@[to_additive (attr := simp)] +theorem coe_mk (f : M ≃* N) (hf1 hf2) : ⇑(mk f hf1 hf2) = f := rfl + +@[to_additive] +theorem toEquiv_eq_coe (f : M ≃ₜ* N) : f.toEquiv = f := + rfl + +@[to_additive (attr := simp)] +theorem toMulEquiv_eq_coe (f : M ≃ₜ* N) : f.toMulEquiv = f := + rfl + +@[to_additive] +theorem toHomeomorph_eq_coe (f : M ≃ₜ* N) : f.toHomeomorph = f := + rfl + +/-- Makes a continuous multiplicative isomorphism from +a homeomorphism which preserves multiplication. -/ +@[to_additive "Makes an continuous additive isomorphism from +a homeomorphism which preserves addition."] +def mk' (f : M ≃ₜ N) (h : ∀ x y, f (x * y) = f x * f y) : M ≃ₜ* N := + ⟨⟨f.toEquiv,h⟩, f.continuous_toFun, f.continuous_invFun⟩ + +set_option linter.docPrime false in -- This is about `ContinuousMulEquiv.mk'` +@[simp] +lemma coe_mk' (f : M ≃ₜ N) (h : ∀ x y, f (x * y) = f x * f y) : ⇑(mk' f h) = f := rfl + +end coe + +section bijective + +@[to_additive] +protected theorem bijective (e : M ≃ₜ* N) : Function.Bijective e := + EquivLike.bijective e + +@[to_additive] +protected theorem injective (e : M ≃ₜ* N) : Function.Injective e := + EquivLike.injective e + +@[to_additive] +protected theorem surjective (e : M ≃ₜ* N) : Function.Surjective e := + EquivLike.surjective e + +@[to_additive] +theorem apply_eq_iff_eq (e : M ≃ₜ* N) {x y : M} : e x = e y ↔ x = y := + e.injective.eq_iff + +end bijective + +section refl + +variable (M) + +/-- The identity map is a continuous multiplicative isomorphism. -/ +@[to_additive (attr := refl) "The identity map is a continuous additive isomorphism."] +def refl : M ≃ₜ* M := + { MulEquiv.refl _ with } + +@[to_additive] +instance : Inhabited (M ≃ₜ* M) := ⟨ContinuousMulEquiv.refl M⟩ + +@[to_additive (attr := simp, norm_cast)] +theorem coe_refl : ↑(refl M) = id := rfl + +@[to_additive (attr := simp)] +theorem refl_apply (m : M) : refl M m = m := rfl + +end refl + +section symm + +/-- The inverse of a ContinuousMulEquiv. -/ +@[to_additive (attr := symm) "The inverse of a ContinuousAddEquiv."] +def symm (cme : M ≃ₜ* N) : N ≃ₜ* M := + { cme.toMulEquiv.symm with + continuous_toFun := cme.continuous_invFun + continuous_invFun := cme.continuous_toFun } +initialize_simps_projections ContinuousMulEquiv (toFun → apply, invFun → symm_apply) + +@[to_additive] +theorem invFun_eq_symm {f : M ≃ₜ* N} : f.invFun = f.symm := rfl + +@[to_additive (attr := simp)] +theorem coe_toHomeomorph_symm (f : M ≃ₜ* N) : (f : M ≃ₜ N).symm = (f.symm : N ≃ₜ M) := rfl + +@[to_additive (attr := simp)] +theorem equivLike_inv_eq_symm (f : M ≃ₜ* N) : EquivLike.inv f = f.symm := rfl + +@[to_additive (attr := simp)] +theorem symm_symm (f : M ≃ₜ* N) : f.symm.symm = f := rfl + +/-- `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`."] +theorem apply_symm_apply (e : M ≃ₜ* N) (y : N) : e (e.symm y) = y := + e.toEquiv.apply_symm_apply y + +/-- `e.symm` is a left inverse of `e`, written as `e.symm (e y) = y`. -/ +@[to_additive (attr := simp) "`e.symm` is a left inverse of `e`, written as `e.symm (e y) = y`."] +theorem symm_apply_apply (e : M ≃ₜ* N) (x : M) : e.symm (e x) = x := + e.toEquiv.symm_apply_apply x + +@[to_additive (attr := simp)] +theorem symm_comp_self (e : M ≃ₜ* N) : e.symm ∘ e = id := + funext e.symm_apply_apply + +@[to_additive (attr := simp)] +theorem self_comp_symm (e : M ≃ₜ* N) : e ∘ e.symm = id := + funext e.apply_symm_apply + +@[to_additive] +theorem apply_eq_iff_symm_apply (e : M ≃ₜ* N) {x : M} {y : N} : e x = y ↔ x = e.symm y := + e.toEquiv.apply_eq_iff_eq_symm_apply + +@[to_additive] +theorem symm_apply_eq (e : M ≃ₜ* N) {x y} : e.symm x = y ↔ x = e y := + e.toEquiv.symm_apply_eq + +@[to_additive] +theorem eq_symm_apply (e : M ≃ₜ* N) {x y} : y = e.symm x ↔ e y = x := + e.toEquiv.eq_symm_apply + +@[to_additive] +theorem eq_comp_symm {α : Type*} (e : M ≃ₜ* N) (f : N → α) (g : M → α) : + f = g ∘ e.symm ↔ f ∘ e = g := + e.toEquiv.eq_comp_symm f g + +@[to_additive] +theorem comp_symm_eq {α : Type*} (e : M ≃ₜ* N) (f : N → α) (g : M → α) : + g ∘ e.symm = f ↔ g = f ∘ e := + e.toEquiv.comp_symm_eq f g + +@[to_additive] +theorem eq_symm_comp {α : Type*} (e : M ≃ₜ* N) (f : α → M) (g : α → N) : + f = e.symm ∘ g ↔ e ∘ f = g := + e.toEquiv.eq_symm_comp f g + +@[to_additive] +theorem symm_comp_eq {α : Type*} (e : M ≃ₜ* N) (f : α → M) (g : α → N) : + e.symm ∘ g = f ↔ g = e ∘ f := + e.toEquiv.symm_comp_eq f g + +end symm + +section trans + +variable {L : Type*} [Mul L] [TopologicalSpace L] + +/-- The composition of two ContinuousMulEquiv. -/ +@[to_additive "The composition of two ContinuousAddEquiv."] +def trans (cme1 : M ≃ₜ* N) (cme2 : N ≃ₜ* L) : M ≃ₜ* L := + { cme1.toMulEquiv.trans cme2.toMulEquiv with + continuous_toFun := by convert Continuous.comp cme2.continuous_toFun cme1.continuous_toFun + continuous_invFun := by convert Continuous.comp cme1.continuous_invFun cme2.continuous_invFun } + +@[to_additive (attr := simp)] +theorem coe_trans (e₁ : M ≃ₜ* N) (e₂ : N ≃ₜ* L) : ↑(e₁.trans e₂) = e₂ ∘ e₁ := rfl + +@[to_additive (attr := simp)] +theorem trans_apply (e₁ : M ≃ₜ* N) (e₂ : N ≃ₜ* L) (m : M) : e₁.trans e₂ m = e₂ (e₁ m) := rfl + +@[to_additive (attr := simp)] +theorem symm_trans_apply (e₁ : M ≃ₜ* N) (e₂ : N ≃ₜ* L) (l : L) : + (e₁.trans e₂).symm l = e₁.symm (e₂.symm l) := rfl + +@[to_additive (attr := simp)] +theorem symm_trans_self (e : M ≃ₜ* N) : e.symm.trans e = refl N := + DFunLike.ext _ _ e.apply_symm_apply + +@[to_additive (attr := simp)] +theorem self_trans_symm (e : M ≃ₜ* N) : e.trans e.symm = refl M := + DFunLike.ext _ _ e.symm_apply_apply + +end trans + +section unique + +/-- The `MulEquiv` between two monoids with a unique element. -/ +@[to_additive "The `AddEquiv` between two `AddMonoid`s with a unique element."] +def ofUnique {M N} [Unique M] [Unique N] [Mul M] [Mul N] + [TopologicalSpace M] [TopologicalSpace N] : M ≃ₜ* N := + { MulEquiv.ofUnique with + continuous_toFun := by continuity + continuous_invFun := by continuity } + +/-- There is a unique monoid homomorphism between two monoids with a unique element. -/ +@[to_additive "There is a unique additive monoid homomorphism between two additive monoids with + a unique element."] +instance {M N} [Unique M] [Unique N] [Mul M] [Mul N] + [TopologicalSpace M] [TopologicalSpace N] : Unique (M ≃ₜ* N) where + default := ofUnique + uniq _ := ext fun _ ↦ Subsingleton.elim _ _ + +end unique + +end ContinuousMulEquiv + +end + +end From e518dee122a6bf6cf7dd7d26e8e9e7873b85089d Mon Sep 17 00:00:00 2001 From: grunweg Date: Thu, 2 Jan 2025 11:03:38 +0000 Subject: [PATCH 067/189] doc(1000.yaml): add some formalisations (#20113) Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- Mathlib/FieldTheory/PrimitiveElement.lean | 6 +- Mathlib/GroupTheory/Coset/Card.lean | 4 + Mathlib/GroupTheory/GroupAction/Quotient.lean | 8 +- Mathlib/GroupTheory/SchurZassenhaus.lean | 4 +- Mathlib/LinearAlgebra/QuadraticForm/Real.lean | 2 +- Mathlib/MeasureTheory/Function/Egorov.lean | 4 +- docs/1000.yaml | 122 ++++++++++++++++++ 7 files changed, 139 insertions(+), 11 deletions(-) diff --git a/Mathlib/FieldTheory/PrimitiveElement.lean b/Mathlib/FieldTheory/PrimitiveElement.lean index b63815fb60c15..6b12f4601768f 100644 --- a/Mathlib/FieldTheory/PrimitiveElement.lean +++ b/Mathlib/FieldTheory/PrimitiveElement.lean @@ -14,10 +14,10 @@ In this file we prove the primitive element theorem. ## Main results -- `exists_primitive_element`: a finite separable extension `E / F` has a primitive element, i.e. - there is an `α : E` such that `F⟮α⟯ = (⊤ : Subalgebra F E)`. +- `Field.exists_primitive_element`: a finite separable extension `E / F` has a primitive element, + i.e. there is an `α : E` such that `F⟮α⟯ = (⊤ : Subalgebra F E)`. -- `exists_primitive_element_iff_finite_intermediateField`: a finite extension `E / F` has a +- `Field.exists_primitive_element_iff_finite_intermediateField`: a finite extension `E / F` has a primitive element if and only if there exist only finitely many intermediate fields between `E` and `F`. diff --git a/Mathlib/GroupTheory/Coset/Card.lean b/Mathlib/GroupTheory/Coset/Card.lean index df9e2e1b9a667..4a76b7a706ad4 100644 --- a/Mathlib/GroupTheory/Coset/Card.lean +++ b/Mathlib/GroupTheory/Coset/Card.lean @@ -8,6 +8,10 @@ import Mathlib.SetTheory.Cardinal.Finite /-! # Lagrange's theorem: the order of a subgroup divides the order of the group. + +* `Subgroup.card_subgroup_dvd_card`: Lagrange's theorem (for multiplicative groups); +there is an analogous version for additive groups + -/ open scoped Pointwise diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index b423ca4fb7406..cdc8548a66f2c 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -17,9 +17,11 @@ import Mathlib.GroupTheory.GroupAction.Hom # Properties of group actions involving quotient groups This file proves properties of group actions which use the quotient group construction, notably -* the orbit-stabilizer theorem `card_orbit_mul_card_stabilizer_eq_card_group` -* the class formula `card_eq_sum_card_group_div_card_stabilizer'` -* Burnside's lemma `sum_card_fixedBy_eq_card_orbits_mul_card_group` +* the orbit-stabilizer theorem `MulAction.card_orbit_mul_card_stabilizer_eq_card_group` +* the class formula `MulAction.card_eq_sum_card_group_div_card_stabilizer'` +* Burnside's lemma `MulAction.sum_card_fixedBy_eq_card_orbits_mul_card_group`, + +as well as their analogues for additive groups. -/ universe u v w diff --git a/Mathlib/GroupTheory/SchurZassenhaus.lean b/Mathlib/GroupTheory/SchurZassenhaus.lean index 84f59877e3187..2e959aacc20e4 100644 --- a/Mathlib/GroupTheory/SchurZassenhaus.lean +++ b/Mathlib/GroupTheory/SchurZassenhaus.lean @@ -12,10 +12,10 @@ In this file we prove the Schur-Zassenhaus theorem. ## Main results -- `exists_right_complement'_of_coprime`: The **Schur-Zassenhaus** theorem: +- `Subgroup.exists_right_complement'_of_coprime`: The **Schur-Zassenhaus** theorem: If `H : Subgroup G` is normal and has order coprime to its index, then there exists a subgroup `K` which is a (right) complement of `H`. -- `exists_left_complement'_of_coprime`: The **Schur-Zassenhaus** theorem: +- `Subgroup.exists_left_complement'_of_coprime`: The **Schur-Zassenhaus** theorem: If `H : Subgroup G` is normal and has order coprime to its index, then there exists a subgroup `K` which is a (left) complement of `H`. -/ diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Real.lean b/Mathlib/LinearAlgebra/QuadraticForm/Real.lean index 69811abb31c01..f330095c26968 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Real.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Real.lean @@ -16,7 +16,7 @@ A real quadratic form is equivalent to a weighted sum of squares with the weights being ±1 or 0. When the real quadratic form is nondegenerate we can take the weights to be ±1, -as in `equivalent_one_zero_neg_one_weighted_sum_squared`. +as in `QuadraticForm.equivalent_one_zero_neg_one_weighted_sum_squared`. -/ diff --git a/Mathlib/MeasureTheory/Function/Egorov.lean b/Mathlib/MeasureTheory/Function/Egorov.lean index ef90331f88c4f..9e40665728a34 100644 --- a/Mathlib/MeasureTheory/Function/Egorov.lean +++ b/Mathlib/MeasureTheory/Function/Egorov.lean @@ -15,8 +15,8 @@ convergence in measure. ## Main results -* `MeasureTheory.Egorov`: Egorov's theorem which shows that a sequence of almost everywhere - convergent functions converges uniformly except on an arbitrarily small set. +* `MeasureTheory.tendstoUniformlyOn_of_ae_tendsto`: Egorov's theorem which shows that a sequence of + almost everywhere convergent functions converges uniformly except on an arbitrarily small set. -/ diff --git a/docs/1000.yaml b/docs/1000.yaml index 08b01979d53fe..27d543c4435d4 100644 --- a/docs/1000.yaml +++ b/docs/1000.yaml @@ -18,6 +18,8 @@ # - perhaps add a wikidata version of the stacks attribute, and generate this file automatically # (can this handle formalisation of *statements* also?) # - complete the information which theorems are already formalised +# +# TODO: display a less ambiguous title for e.g. "Liouville's theorem" Q11518: title: Pythagorean theorem @@ -85,6 +87,7 @@ Q180345: Q182505: title: Bayes' theorem + decl: ProbabilityTheory.cond_eq_inv_mul_cond_mul Q184410: title: Four color theorem @@ -98,6 +101,8 @@ Q188745: Q189136: title: Mean value theorem decl: exists_deriv_eq_slope + # XXX: other branch found exists_ratio_deriv_eq_ratio_slope, + # is about Cauchy's vs Lagrange's Mean Value Theorem, TODO decide! author: Yury G. Kudryashov Q190026: @@ -123,9 +128,11 @@ Q192760: Q193286: title: Rolle's theorem + decl: exists_deriv_eq_zero Q193878: title: Chinese remainder theorem + decl: Ideal.quotientInfRingEquivPiQuotient Q193882: title: Menelaus's theorem @@ -182,6 +189,8 @@ Q213603: Q220680: title: Banach fixed-point theorem + ## Contraction mapping theorem in mathlib + decl: ContractingWith.exists_fixedPoint Q225973: title: Ore's theorem @@ -208,6 +217,7 @@ Q245098: Q245098X: title: Bolzano's theorem + # variant of the intermediate value theorem above, TODO should exist in mathlib, add decl Q245902: title: Riesz–Thorin theorem @@ -217,6 +227,9 @@ Q248931: Q253214: title: Heine–Borel theorem + # TODO: two versions in mathlib, which one do we want? + # decl: Metric.isCompact_iff_isClosed_bounded + # decl: FiniteDimensional.proper Q255996: title: Equipartition theorem @@ -273,12 +286,18 @@ Q287347: Q288465: title: Orbit-stabilizer theorem + decl: MulAction.orbitEquivQuotientStabilizer Q303402: title: Rank–nullity theorem + # TODO: this is also in mathlib, TODO add decl (and mention in overview) Q318767: title: Abel's theorem + decls: + - Complex.tendsto_tsum_powerSeries_nhdsWithin_stolzCone + - Real.tendsto_tsum_powerSeries_nhdsWithin_lt + author: Jeremy Tan Q321237: title: Green's theorem @@ -330,21 +349,27 @@ Q384142: Q386292: title: Prime number theorem + # also in 100.yaml Q388525: title: Bell's theorem Q401319: title: Puiseux's theorem + # also in 100.yaml Q401905: title: Skorokhod's representation theorem Q420714: title: Akra–Bazzi theorem + decl: AkraBazziRecurrence.isBigO_asympBound Q422187: title: Myhill–Nerode theorem + # two in-progress formalisations, none merged as of end of 2024 + # https://github.com/atarnoam/lean-automata/blob/main/src/regular_languages.lean + # https://github.com/leanprover-community/mathlib4/pull/11311 Q425432: title: Laurent expansion theorem @@ -365,6 +390,9 @@ Q467756: Q468391: title: Bolzano–Weierstrass theorem + decls: + - tendsto_subseq_of_frequently_bounded # less standard version + - tendsto_subseq_of_bounded Q470877: title: Stirling's theorem @@ -403,6 +431,8 @@ Q504843: Q505798: title: Lagrange's theorem + # group theorem result; exist additive and multiplicative version + decl: Subgroup.card_subgroup_dvd_card Q510197: title: Tutte theorem @@ -424,15 +454,18 @@ Q528987: Q530152: title: Picard–Lindelöf theorem + decl: IsPicardLindelof.exists_forall_hasDerivWithinAt_Icc_eq Q535366: title: Montel's theorem Q536640: title: Hall's marriage theorem + decl: Finset.all_card_le_biUnion_card_iff_exists_injective Q537618: title: Banach–Alaoglu theorem + decl: WeakDual.isCompact_polar Q544292: title: Lagrange inversion theorem @@ -453,12 +486,14 @@ Q567843: Q570779: title: Vieta's formulas + decl: Multiset.prod_X_add_C_eq_sum_esymm Q574902: title: Tarski's indefinability theorem Q576478: title: Liouville's theorem + decl: Differentiable.apply_eq_apply_of_bounded Q578555: title: Noether's theorem @@ -468,6 +503,8 @@ Q579515: Q583147: title: Sard's theorem + # exists a statement in Lean, WIP formalisations in https://github.com/urkud/SardMoreira + # and https://github.com/fpvandoorn/sard (manifold version) Q586051: title: Haag–Łopuszański–Sohnius theorem @@ -501,12 +538,15 @@ Q617417: Q619985: title: Multinomial theorem + decl: Finset.sum_pow_eq_sum_piAntidiag_of_commute + # TODO: is Finset.sum_pow_of_commute better? Q620602: title: Virial theorem Q631561: title: Closed range theorem + # TODO: this surely exists in mathlib, add decl! Q632546X: title: Sylvester's theorem @@ -521,6 +561,8 @@ Q637418: Q642620: title: Krein–Milman theorem + decl: closure_convexHull_extremePoints + author: Yaël Dillies Q643513: title: Riesz–Fischer theorem @@ -534,6 +576,7 @@ Q646523: Q649469: title: Modularity theorem + # far away; FLT project is working on a special case Q649977: title: Shirshov–Cohn theorem @@ -549,9 +592,11 @@ Q656176: Q656198: title: Maschke's theorem + # `RepresentationTheory/Maschke` comes very close, but doesn't prove the standard form yet... Q656645: title: Hilbert's basis theorem + # search mathlib or zulip; mathlib is close to this! Q656772: title: Cayley–Hamilton theorem @@ -574,6 +619,8 @@ Q657903: Q660799: title: Darboux's theorem + # "all derivatives have the intermediate value property" + # TODO should exist in mathlib, add decl! Q670235: title: Fundamental theorem of arithmetic @@ -606,6 +653,8 @@ Q693083: Q716171: title: Erdős–Ginzburg–Ziv theorem + decl: ZMod.erdos_ginzburg_ziv + author: Yaël Dillies Q718875: title: Erdős–Ko–Rado theorem @@ -630,9 +679,11 @@ Q730222: Q733081: title: Impossibility of angle trisection + # first part of problem 8 in 100.yaml Q737851: title: Banach–Tarski theorem + # WIP formalisation at https://github.com/Bergschaf/banach-tarski Q737892: title: Bendixson–Dulac theorem @@ -657,6 +708,7 @@ Q751120: Q752375: title: Extreme value theorem + decl: IsCompact.exists_isMinOn Q755986: title: Atiyah–Bott fixed-point theorem @@ -671,15 +723,21 @@ Q756946: Q764287: title: Van der Waerden's theorem + decl: Combinatorics.exists_mono_homothetic_copy + author: David Wärn + # TODO: this is only some version; should this entry be modified accordingly? Q765987: title: Heine–Cantor theorem + # XXX: does mathlib have this? Q766722: title: Liouville's theorem + # theorem about Hamiltonian mechanics Q776578: title: Artin–Wedderburn theorem + # WIP formalisation, but not complete yet (Dec 1, 2024) Q777924: title: Tijdeman's theorem @@ -750,6 +808,7 @@ Q848092: Q848375: title: Implicit function theorem + decl: ImplicitFunctionData.implicitFunction Q848810: title: Kronecker's theorem @@ -794,6 +853,7 @@ Q865665: Q866116: title: Hahn–Banach theorem + decl: exists_extension_norm_eq Q867141: title: Ehresmann's theorem @@ -880,6 +940,7 @@ Q929547: Q931001: title: Inverse function theorem + decl: HasStrictFDerivAt.to_localInverse Q931001X: title: Constant rank theorem @@ -898,12 +959,16 @@ Q942046: Q943246: title: Wedderburn's little theorem + decl: littleWedderburn Q944297: title: Open mapping theorem + # TODO: this entry points to Wikipedia's disambiguation page => wikipedia should be fixed + # Banach open mapping theorem/Banach-Schauder theorem in functional analysis is in mathlib (ContinuousLinearMap.isOpenMap) Q948664: title: Kneser's theorem + # TODO: is this in mathlib already? Q951327: title: Pasch's theorem @@ -922,6 +987,7 @@ Q966837: Q967972: title: Open mapping theorem + # result in complex analysis. TODO does mathlib have this? Q973359: title: Rado's theorem @@ -960,6 +1026,8 @@ Q1008566: Q1032886: title: Hales–Jewett theorem + decl: Combinatorics.Line.exists_mono_in_high_dimension + author: David Wärn Q1033910: title: Cantor–Bernstein–Schroeder theorem @@ -971,12 +1039,15 @@ Q1037559: Q1038716: title: Identity theorem + # complex analysis, holomorphic functions: TODO add decl! Q1046232: title: Szemerédi's theorem Q1047749: title: Turán's theorem + decl: SimpleGraph.isTuranMaximal_iff_nonempty_iso_turanGraph + author: Jeremy Tan Q1048589: title: Hohenberg–Kohn theorems @@ -989,9 +1060,14 @@ Q1050037: Q1050203: title: Cantor's intersection theorem + # many versions; are this (and friends, in the same file) all of the wikipedia article? + decl: IsCompact.nonempty_sInter_of_directed_nonempty_isCompact_isClosed Q1050932: title: Hartogs's theorem + url: https://github.com/girving/ray/blob/main/Ray/Hartogs/Hartogs.lean + identifiers: Pair.hartogs + author: Geoffrey Irving Q1051404: title: Cesàro's theorem @@ -1001,6 +1077,7 @@ Q1052021: Q1052678: title: Baire category theorem + decl: dense_sInter_of_isOpen Q1052752: title: Stolz–Cesàro theorem @@ -1021,6 +1098,7 @@ Q1057919: Q1058662: title: Darboux's theorem + # symplectic geometry result Q1059151: title: FWL theorem @@ -1030,27 +1108,38 @@ Q1064471: Q1065257: title: Squeeze theorem + decl: tendsto_of_tendsto_of_tendsto_of_le_of_le' Q1065966: title: Isomorphism theorem + # Three isomorphism theorems in abstract algebra + decls: + - QuotientGroup.quotientKerEquivRange + - QuotientGroup.quotientInfEquivProdNormalQuotient + - QuotientGroup.quotientQuotientEquivQuotient Q1067156: title: Dominated convergence theorem + decl: MeasureTheory.tendsto_integral_of_dominated_convergence Q1067156X: title: Bounded convergence theorem + # TODO: is this in mathlib already? Q1068085: title: Closed graph theorem + decl: LinearMap.continuous_of_isClosed_graph Q1068283: title: Löwenheim–Skolem theorem + decl: FirstOrder.Language.exists_elementaryEmbedding_card_eq Q1068385: title: Clairaut's theorem Q1068976: title: Hilbert's Nullstellensatz + decl: MvPolynomial.vanishingIdeal_zeroLocus_eq_radical Q1075398: title: Jung's theorem @@ -1071,6 +1160,8 @@ Q1082910: Q1095330: title: Apéry's theorem + url: https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/pi.20is.20irrational.3F/near/478674074 + # TODO: insert a link to a repository, and perhaps get that into mathlib :-) Q1097021: title: Minkowski's theorem @@ -1118,6 +1209,7 @@ Q1136262: Q1137014: title: Tychonoff's theorem + decl: isCompact_pi_infinite Q1137206: title: Taylor's theorem @@ -1128,6 +1220,8 @@ Q1137206: Q1139041: title: Cauchy's theorem + # for finite groups + decl: exists_prime_orderOf_dvd_card Q1139430: title: Hurwitz's theorem @@ -1158,9 +1252,11 @@ Q1146791: Q1148215: title: Mitchell's embedding theorem + # there was a project talking about this; very in progress, I believe. statement exists. Q1149022: title: Fubini's theorem + decl: MeasureTheory.integral_prod Q1149185X: title: Kirby–Paris theorem @@ -1170,6 +1266,7 @@ Q1149185: Q1149458: title: Compactness theorem + decl: FirstOrder.Language.Theory.isSatisfiable_iff_isFinitelySatisfiable Q1149522: title: Lami's theorem @@ -1182,6 +1279,9 @@ Q1153441: Q1153584: title: Monotone convergence theorem + # TODO: many versions in mathlib, choose one + # e.g. integral_tendsto_of_tendsto_of_antitone (Bochner integral) + # or lintegral_tendsto_of_tendsto_of_monotone (Lebesgue integral) Q1166774: title: Stone's representation theorem for Boolean algebras @@ -1200,6 +1300,7 @@ Q1186134: Q1186808: title: Lucas's theorem + decl: Choose.lucas_theorem Q1187646: title: Fundamental theorem on homomorphisms @@ -1209,21 +1310,27 @@ Q1188048: Q1188392: title: Zeckendorf's theorem + decl: Nat.zeckendorfEquiv Q1188504: title: Pitman–Koopman–Darmois theorem Q1191319: title: Radon–Nikodym theorem + decl: MeasureTheory.Measure.absolutelyContinuous_iff_withDensity_rnDeriv_eq Q1191709: title: Egorov's theorem + decl: MeasureTheory.tendstoUniformlyOn_of_ae_tendsto + author: Kexing Ying Q1191862: title: Rouché's theorem Q1194053: title: Metrization theorems + decl: TopologicalSpace.metrizableSpace_of_t3_secondCountable + comment: "Urysohn's metrization theorem (only)" Q1196538: title: Descartes's theorem @@ -1264,6 +1371,7 @@ Q1249069: Q1259814: title: Nielsen–Schreier theorem + decl: subgroupIsFreeOfIsFree Q1276318: title: Schwartz kernel theorem @@ -1276,6 +1384,8 @@ Q1306092: Q1306095: title: Whitney embedding theorem + decl: exists_embedding_euclidean_of_compact + comment: "baby version: for compact manifolds; embedding into some *n*" Q1307676: title: Künneth theorem @@ -1303,6 +1413,7 @@ Q1340623: Q1346677: title: Tietze extension theorem + decl: BoundedContinuousFunction.exists_extension_norm_eq_of_isClosedEmbedding Q1347094: title: Alternate Interior Angles Theorem @@ -1360,6 +1471,7 @@ Q1425529: Q1426292: title: Banach–Steinhaus theorem + decl: banach_steinhaus Q1434158: title: Fluctuation dissipation theorem @@ -1387,6 +1499,7 @@ Q1472120: Q1477053: title: Arzelà–Ascoli theorem + decl: BoundedContinuousFunction.arzela_ascoli Q1505529: title: Runge's theorem @@ -1427,6 +1540,8 @@ Q1564541: Q1566341: title: Hindman's theorem + decl: Hindman.FP_partition_regular + author: David Wärn Q1566372: title: Haag's theorem @@ -1436,9 +1551,11 @@ Q1568612: Q1568811: title: Hahn decomposition theorem + decl: MeasureTheory.SignedMeasure.exists_isCompl_positive_negative Q1572474: title: Lindemann–Weierstrass theorem + # WIP PR: https://github.com/leanprover-community/mathlib4/pull/6718 Q1576235: title: Lochs's theorem @@ -1517,6 +1634,7 @@ Q1752516: Q1752621: title: Sylvester's law of inertia + decl: QuadraticForm.equivalent_one_zero_neg_one_weighted_sum_squared Q1752988: title: Kutta–Joukowski theorem @@ -1634,6 +1752,7 @@ Q2071632: Q2093886: title: Primitive element theorem + decl: Field.exists_primitive_element Q2094241: title: Hopf–Rinow theorem @@ -1710,6 +1829,7 @@ Q2226786: Q2226800: title: Schur–Zassenhaus theorem + decl: Subgroup.exists_left_complement'_of_coprime Q2226803: title: Tennenbaum's theorem @@ -3367,6 +3487,7 @@ Q7818977: Q7820874: title: Tonelli's theorem + # TODO: this is on mathlib, right? Q7824894: title: Topkis's theorem @@ -3379,6 +3500,7 @@ Q7825663: Q7827204: title: Mazur's torsion theorem + # TODO: is this used in FLT? Q7841060: title: Trichotomy theorem From 45302a80e5099717640847a191d1f5f02c9acb44 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 2 Jan 2025 11:03:40 +0000 Subject: [PATCH 068/189] feat: functoriality of `Measure.comap` (#20218) --- Mathlib/MeasureTheory/Measure/Comap.lean | 22 ++++++++++++++++++++-- 1 file changed, 20 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/Comap.lean b/Mathlib/MeasureTheory/Measure/Comap.lean index adefd18c588f4..4b7a871e2f74c 100644 --- a/Mathlib/MeasureTheory/Measure/Comap.lean +++ b/Mathlib/MeasureTheory/Measure/Comap.lean @@ -26,7 +26,7 @@ namespace MeasureTheory namespace Measure -variable {α β : Type*} {s : Set α} +variable {α β γ : Type*} {s : Set α} open Classical in /-- Pullback of a `Measure` as a linear map. If `f` sends each measurable set to a measurable @@ -62,7 +62,8 @@ def comap [MeasurableSpace α] [MeasurableSpace β] (f : α → β) (μ : Measur exact (measure_inter_add_diff₀ _ (hf.2 s hs)).symm else 0 -variable {ma : MeasurableSpace α} {mb : MeasurableSpace β} +variable {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} + {f : α → β} {g : β → γ} theorem comap_apply₀ (f : α → β) (μ : Measure β) (hfi : Injective f) (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) @@ -129,6 +130,23 @@ theorem comap_preimage (f : α → β) (μ : Measure β) (hf : Injective f) (hf' · simp [comap, hf] · simp [comap, hf] +@[simp] +lemma comap_id (μ : Measure β) : comap (fun x ↦ x) μ = μ := by + ext s hs + rw [comap_apply, image_id'] + · exact injective_id + all_goals simp [*] + +lemma comap_comap (hf' : ∀ s, MeasurableSet s → MeasurableSet (f '' s)) (hg : Injective g) + (hg' : ∀ s, MeasurableSet s → MeasurableSet (g '' s)) (μ : Measure γ) : + comap f (comap g μ) = comap (g ∘ f) μ := by + by_cases hf : Injective f + · ext s hs + rw [comap_apply _ hf hf' _ hs, comap_apply _ hg hg' _ (hf' _ hs), + comap_apply _ (hg.comp hf) (fun t ht ↦ image_comp g f _ ▸ hg' _ <| hf' _ ht) _ hs, image_comp] + · rw [comap, dif_neg <| mt And.left hf, comap, dif_neg fun h ↦ hf h.1.of_comp] + + end Measure end MeasureTheory From 0673d5a7da4637c09c17d6f66eb562e15b60f0d2 Mon Sep 17 00:00:00 2001 From: Bergschaf Date: Thu, 2 Jan 2025 11:03:41 +0000 Subject: [PATCH 069/189] feat(Order/CompleteLattice): instances for CompleteSemilatticeInf / CompleteSemilatticeSup for Dual types (#20246) --- Mathlib/Order/CompleteLattice.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index edf6f099fdb2b..c225c90223d8d 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -173,6 +173,15 @@ theorem sInf_singleton {a : α} : sInf {a} = a := end +instance {α : Type*} [CompleteSemilatticeInf α] : CompleteSemilatticeSup αᵒᵈ where + le_sSup := @CompleteSemilatticeInf.sInf_le α _ + sSup_le := @CompleteSemilatticeInf.le_sInf α _ + +instance {α : Type*} [CompleteSemilatticeSup α] : CompleteSemilatticeInf αᵒᵈ where + le_sInf := @CompleteSemilatticeSup.sSup_le α _ + sInf_le := @CompleteSemilatticeSup.le_sSup α _ + + /-- A complete lattice is a bounded lattice which has suprema and infima for every subset. -/ class CompleteLattice (α : Type*) extends Lattice α, CompleteSemilatticeSup α, CompleteSemilatticeInf α, Top α, Bot α where From ff2c9d6d2d2cdccc9bf94b4f246b44aa71e6e338 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 2 Jan 2025 11:03:42 +0000 Subject: [PATCH 070/189] feat(BooleanSubalgebra): closure of supClosed+infClosed set containing bot+top (#20251) Every element in the closure of such a set `s` can be written as a finite sup of elements of the form `a \ b` with `a`, `b` in `s`. --- Mathlib/Order/BooleanSubalgebra.lean | 38 ++++++++++++++++++++++++++++ 1 file changed, 38 insertions(+) diff --git a/Mathlib/Order/BooleanSubalgebra.lean b/Mathlib/Order/BooleanSubalgebra.lean index ca4126280f883..24751706b1099 100644 --- a/Mathlib/Order/BooleanSubalgebra.lean +++ b/Mathlib/Order/BooleanSubalgebra.lean @@ -361,6 +361,44 @@ lemma closure_bot_sup_induction {p : ∀ g ∈ closure s, Prop} (mem : ∀ x hx, compl_mem' := fun ⟨_, hb⟩ ↦ ⟨_, compl _ _ hb⟩ } closure_le (L := L).mpr (fun y hy ↦ ⟨subset_closure hy, mem y hy⟩) hx |>.elim fun _ ↦ id +section sdiff_sup + +variable (sup : SupClosed s) (inf : InfClosed s) (bot_mem : ⊥ ∈ s) (top_mem : ⊤ ∈ s) +include sup inf bot_mem top_mem + +theorem mem_closure_iff_sup_sdiff {a : α} : + a ∈ closure s ↔ ∃ t : Finset (s × s), a = t.sup fun x ↦ x.1.1 \ x.2.1 := by + classical + refine ⟨closure_bot_sup_induction + (fun x h ↦ ⟨{(⟨x, h⟩, ⟨⊥, bot_mem⟩)}, by simp⟩) ⟨∅, by simp⟩ ?_ ?_, ?_⟩ + · rintro ⟨t, rfl⟩ + exact t.sup_mem _ (subset_closure bot_mem) (fun _ h _ ↦ sup_mem h) _ + fun x hx ↦ sdiff_mem (subset_closure x.1.2) (subset_closure x.2.2) + · rintro _ - _ - ⟨t₁, rfl⟩ ⟨t₂, rfl⟩ + exact ⟨t₁ ∪ t₂, by rw [Finset.sup_union]⟩ + rintro x - ⟨t, rfl⟩ + refine t.induction ⟨{(⟨⊤, top_mem⟩, ⟨⊥, bot_mem⟩)}, by simp⟩ fun ⟨x, y⟩ t _ ⟨tc, eq⟩ ↦ ?_ + simp_rw [Finset.sup_insert, compl_sup, eq] + refine tc.induction ⟨∅, by simp⟩ fun ⟨z, w⟩ tc _ ⟨t, eq⟩ ↦ ?_ + simp_rw [Finset.sup_insert, inf_sup_left, eq] + refine ⟨{(z, ⟨_, sup x.2 w.2⟩), (⟨_, inf y.2 z.2⟩, w)} ∪ t, ?_⟩ + simp_rw [Finset.sup_union, Finset.sup_insert, Finset.sup_singleton, sdiff_eq, + compl_sup, inf_left_comm z.1, compl_inf, compl_compl, inf_sup_right, inf_assoc] + +@[elab_as_elim] theorem closure_sdiff_sup_induction {p : ∀ g ∈ closure s, Prop} + (sdiff : ∀ x hx y hy, p (x \ y) (sdiff_mem (subset_closure hx) (subset_closure hy))) + (sup' : ∀ x hx y hy, p x hx → p y hy → p (x ⊔ y) (sup_mem hx hy)) + {x} (hx : x ∈ closure s) : p x hx := by + obtain ⟨t, rfl⟩ := (mem_closure_iff_sup_sdiff sup inf bot_mem top_mem).mp hx + revert hx + classical + refine t.induction (by simpa using sdiff _ bot_mem _ bot_mem) fun x t _ ih hxt ↦ ?_ + simp only [Finset.sup_insert] at hxt ⊢ + exact sup' _ _ _ _ (sdiff _ x.1.2 _ x.2.2) + (ih <| (mem_closure_iff_sup_sdiff sup inf bot_mem top_mem).mpr ⟨_, rfl⟩) + +end sdiff_sup + end BooleanAlgebra section CompleteBooleanAlgebra From 758574f8316e0f3d201f50f85dd82d5bb6b3bc82 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 2 Jan 2025 11:03:43 +0000 Subject: [PATCH 071/189] feat(Algebra): auxiliary lemmas for flatness over semirings (#20265) --- Mathlib/Algebra/DirectSum/Module.lean | 19 ++++++ Mathlib/Algebra/Module/Injective.lean | 6 ++ Mathlib/Data/DFinsupp/Defs.lean | 5 ++ .../DirectSum/TensorProduct.lean | 5 ++ .../LinearAlgebra/TensorProduct/Tower.lean | 5 ++ .../RingTheory/Finiteness/Cardinality.lean | 16 +++-- .../RingTheory/Finiteness/TensorProduct.lean | 66 ++++++++++++------- Mathlib/RingTheory/Flat/CategoryTheory.lean | 4 +- .../RingTheory/Localization/BaseChange.lean | 5 +- 9 files changed, 98 insertions(+), 33 deletions(-) diff --git a/Mathlib/Algebra/DirectSum/Module.lean b/Mathlib/Algebra/DirectSum/Module.lean index dba8c5158d5d0..73ae9c2eaa56f 100644 --- a/Mathlib/Algebra/DirectSum/Module.lean +++ b/Mathlib/Algebra/DirectSum/Module.lean @@ -202,6 +202,25 @@ 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 +section map + +variable {R} {N : ι → Type*} [∀ i, AddCommMonoid (N i)] [∀ i, Module R (N i)] + (f : Π i, M i →ₗ[R] N i) + +/-- The linear map between direct sums induced by a family of linear maps. -/ +def lmap : (⨁ i, M i) →ₗ[R] ⨁ i, N i := DFinsupp.mapRange.linearMap f + +@[simp] theorem lmap_apply (x i) : lmap f x i = f i (x i) := rfl + +@[simp] theorem lmap_lof [DecidableEq ι] (i) (x : M i) : + lmap f (lof R _ _ _ x) = lof R _ _ _ (f i x) := + DFinsupp.mapRange_single (hf := fun _ ↦ map_zero _) + +theorem lmap_injective : Function.Injective (lmap f) ↔ ∀ i, Function.Injective (f i) := by + classical exact DFinsupp.mapRange_injective (hf := fun _ ↦ map_zero _) + +end map + section CongrLeft variable {κ : Type*} diff --git a/Mathlib/Algebra/Module/Injective.lean b/Mathlib/Algebra/Module/Injective.lean index c18c85f2eaa98..845e50bfd140e 100644 --- a/Mathlib/Algebra/Module/Injective.lean +++ b/Mathlib/Algebra/Module/Injective.lean @@ -67,6 +67,12 @@ namespace Module.Baer variable {R Q} {M N : Type*} [AddCommGroup M] [AddCommGroup N] variable [Module R M] [Module R N] (i : M →ₗ[R] N) (f : M →ₗ[R] Q) +lemma of_equiv (e : Q ≃ₗ[R] M) (h : Module.Baer R Q) : Module.Baer R M := fun I g ↦ + have ⟨g', h'⟩ := h I (e.symm ∘ₗ g) + ⟨e ∘ₗ g', by simpa [LinearEquiv.eq_symm_apply] using h'⟩ + +lemma congr (e : Q ≃ₗ[R] M) : Module.Baer R Q ↔ Module.Baer R M := ⟨of_equiv e, of_equiv e.symm⟩ + /-- If we view `M` as a submodule of `N` via the injective linear map `i : M ↪ N`, then a submodule between `M` and `N` is a submodule `N'` of `N`. To prove Baer's criterion, we need to consider pairs of `(N', f')` such that `M ≤ N' ≤ N` and `f'` extends `f`. -/ diff --git a/Mathlib/Data/DFinsupp/Defs.lean b/Mathlib/Data/DFinsupp/Defs.lean index cdb8f19b4f3a3..30ed70de9b564 100644 --- a/Mathlib/Data/DFinsupp/Defs.lean +++ b/Mathlib/Data/DFinsupp/Defs.lean @@ -922,6 +922,11 @@ theorem mapRange_single {f : ∀ i, β₁ i → β₂ i} {hf : ∀ i, f i 0 = 0} simp · simp [h, hf] +theorem mapRange_injective (f : ∀ i, β₁ i → β₂ i) (hf : ∀ i, f i 0 = 0) : + Function.Injective (mapRange f hf) ↔ ∀ i, Function.Injective (f i) := + ⟨fun h i x y eq ↦ single_injective (@h (single i x) (single i y) <| by + simpa using congr_arg _ eq), fun h _ _ eq ↦ DFinsupp.ext fun i ↦ h i congr($eq i)⟩ + variable [∀ (i) (x : β₁ i), Decidable (x ≠ 0)] [∀ (i) (x : β₂ i), Decidable (x ≠ 0)] theorem support_mapRange {f : ∀ i, β₁ i → β₂ i} {hf : ∀ i, f i 0 = 0} {g : Π₀ i, β₁ i} : diff --git a/Mathlib/LinearAlgebra/DirectSum/TensorProduct.lean b/Mathlib/LinearAlgebra/DirectSum/TensorProduct.lean index 1a6c8831e3b9d..e461a07336417 100644 --- a/Mathlib/LinearAlgebra/DirectSum/TensorProduct.lean +++ b/Mathlib/LinearAlgebra/DirectSum/TensorProduct.lean @@ -181,6 +181,11 @@ theorem directSumRight_symm_lof_tmul (x : M₁') (i : ι₂) (y : M₂ i) : x ⊗ₜ[R] DirectSum.lof R _ _ i y := by rw [LinearEquiv.symm_apply_eq, directSumRight_tmul_lof] +lemma directSumRight_comp_rTensor (f : M₁' →ₗ[R] M₂'): + (directSumRight R M₂' M₁).toLinearMap ∘ₗ f.rTensor _ = + (lmap fun _ ↦ f.rTensor _) ∘ₗ directSumRight R M₁' M₁ := by + ext; simp + end TensorProduct end Ring diff --git a/Mathlib/LinearAlgebra/TensorProduct/Tower.lean b/Mathlib/LinearAlgebra/TensorProduct/Tower.lean index 95ba16af73488..b55604d42a282 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Tower.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Tower.lean @@ -437,6 +437,11 @@ theorem cancelBaseChange_symm_tmul (m : M) (n : N) : (cancelBaseChange R A B M N).symm (m ⊗ₜ n) = m ⊗ₜ (1 ⊗ₜ n) := rfl +theorem lTensor_comp_cancelBaseChange (f : N →ₗ[R] Q) : + lTensor _ _ f ∘ₗ cancelBaseChange R A B M N = + (cancelBaseChange R A B M Q).toLinearMap ∘ₗ lTensor _ _ (lTensor _ _ f) := by + ext; simp + end cancelBaseChange section leftComm diff --git a/Mathlib/RingTheory/Finiteness/Cardinality.lean b/Mathlib/RingTheory/Finiteness/Cardinality.lean index e03f6de91613f..721a173413698 100644 --- a/Mathlib/RingTheory/Finiteness/Cardinality.lean +++ b/Mathlib/RingTheory/Finiteness/Cardinality.lean @@ -20,7 +20,8 @@ open Finsupp section ModuleAndAlgebra -variable (R A B M N : Type*) +universe v u +variable (R : Type u) (A B M N : Type*) namespace Module @@ -30,18 +31,23 @@ namespace Finite open Submodule Set -variable {R M N} - -variable (R M) in +/-- A finite module admits a surjective linear map from a finite free module. -/ lemma exists_fin' [Module.Finite R M] : ∃ (n : ℕ) (f : (Fin n → R) →ₗ[R] M), Surjective f := by have ⟨n, s, hs⟩ := exists_fin (R := R) (M := M) refine ⟨n, Basis.constr (Pi.basisFun R _) ℕ s, ?_⟩ rw [← LinearMap.range_eq_top, Basis.constr_range, hs] -variable (R) in +lemma small [Module.Finite R M] [Small.{v} R] : Small.{v} M := + have ⟨_, _, h⟩ := exists_fin' R M + small_of_surjective h + +variable {M} + lemma _root_.Module.finite_of_finite [Finite R] [Module.Finite R M] : Finite M := by obtain ⟨n, f, hf⟩ := exists_fin' R M; exact .of_surjective f hf +variable {R} + @[deprecated (since := "2024-10-13")] alias _root_.FiniteDimensional.finite_of_finite := finite_of_finite diff --git a/Mathlib/RingTheory/Finiteness/TensorProduct.lean b/Mathlib/RingTheory/Finiteness/TensorProduct.lean index 9e513013be264..e1a686b2c3209 100644 --- a/Mathlib/RingTheory/Finiteness/TensorProduct.lean +++ b/Mathlib/RingTheory/Finiteness/TensorProduct.lean @@ -23,35 +23,51 @@ open Finsupp namespace Submodule -variable {R : Type*} {M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] - -open Set - -variable {P : Type*} [AddCommMonoid P] [Module R P] -variable (f : M →ₗ[R] P) +variable {R M N : Type*} [CommSemiring R] [AddCommMonoid M] + [AddCommMonoid N] [Module R M] [Module R N] {I : Submodule R N} + +open TensorProduct LinearMap +/-- Every `x : N ⊗ M` is the image of some `y : J ⊗ M`, where `J` is a finitely generated +submodule of `N`, under the tensor product of the inclusion `J → N` and the identity `M → M`. -/ +theorem exists_fg_le_eq_rTensor_subtype (x : N ⊗ M) : + ∃ (J : Submodule R N) (_ : J.FG) (y : J ⊗ M), x = rTensor M J.subtype y := by + induction x with + | zero => exact ⟨⊥, fg_bot, 0, rfl⟩ + | tmul i m => exact ⟨R ∙ i, fg_span_singleton i, ⟨i, mem_span_singleton_self _⟩ ⊗ₜ[R] m, rfl⟩ + | add x₁ x₂ ihx₁ ihx₂ => + obtain ⟨J₁, fg₁, y₁, rfl⟩ := ihx₁ + obtain ⟨J₂, fg₂, y₂, rfl⟩ := ihx₂ + refine ⟨J₁ ⊔ J₂, fg₁.sup fg₂, + rTensor M (J₁.inclusion le_sup_left) y₁ + rTensor M (J₂.inclusion le_sup_right) y₂, ?_⟩ + rw [map_add, ← rTensor_comp_apply, ← rTensor_comp_apply] + rfl -variable {f} +theorem exists_fg_le_subset_range_rTensor_subtype (s : Set (N ⊗[R] M)) (hs : s.Finite) : + ∃ (J : Submodule R N) (_ : J.FG), s ⊆ LinearMap.range (rTensor M J.subtype) := by + choose J fg y eq using exists_fg_le_eq_rTensor_subtype (R := R) (M := M) (N := N) + rw [← Set.finite_coe_iff] at hs + refine ⟨⨆ x : s, J x, fg_iSup _ fun _ ↦ fg _, fun x hx ↦ + ⟨rTensor M (inclusion <| le_iSup _ ⟨x, hx⟩) (y x), .trans ?_ (eq x).symm⟩⟩ + rw [← comp_apply, ← rTensor_comp]; rfl -open TensorProduct LinearMap in +open TensorProduct LinearMap /-- Every `x : I ⊗ M` is the image of some `y : J ⊗ M`, where `J ≤ I` is finitely generated, under the tensor product of `J.inclusion ‹J ≤ I› : J → I` and the identity `M → M`. -/ -theorem exists_fg_le_eq_rTensor_inclusion {R M N : Type*} [CommRing R] [AddCommGroup M] - [AddCommGroup N] [Module R M] [Module R N] {I : Submodule R N} (x : I ⊗ M) : - ∃ (J : Submodule R N) (_ : J.FG) (hle : J ≤ I) (y : J ⊗ M), - x = rTensor M (J.inclusion hle) y := by - induction x with - | zero => exact ⟨⊥, fg_bot, zero_le _, 0, rfl⟩ - | tmul i m => exact ⟨R ∙ i.val, fg_span_singleton i.val, - (span_singleton_le_iff_mem _ _).mpr i.property, - ⟨i.val, mem_span_singleton_self _⟩ ⊗ₜ[R] m, rfl⟩ - | add x₁ x₂ ihx₁ ihx₂ => - obtain ⟨J₁, hfg₁, hle₁, y₁, rfl⟩ := ihx₁ - obtain ⟨J₂, hfg₂, hle₂, y₂, rfl⟩ := ihx₂ - refine ⟨J₁ ⊔ J₂, hfg₁.sup hfg₂, sup_le hle₁ hle₂, - rTensor M (J₁.inclusion (le_sup_left : J₁ ≤ J₁ ⊔ J₂)) y₁ + - rTensor M (J₂.inclusion (le_sup_right : J₂ ≤ J₁ ⊔ J₂)) y₂, ?_⟩ - rewrite [map_add, ← rTensor_comp_apply, ← rTensor_comp_apply] - rfl +theorem exists_fg_le_eq_rTensor_inclusion (x : I ⊗ M) : + ∃ (J : Submodule R N) (_ : J.FG) (hle : J ≤ I) (y : J ⊗ M), + x = rTensor M (J.inclusion hle) y := by + obtain ⟨J, fg, y, rfl⟩ := exists_fg_le_eq_rTensor_subtype x + refine ⟨J.map I.subtype, fg.map _, I.map_subtype_le J, rTensor M (I.subtype.submoduleMap J) y, ?_⟩ + rw [← LinearMap.rTensor_comp_apply]; rfl + +theorem exists_fg_le_subset_range_rTensor_inclusion (s : Set (I ⊗[R] M)) (hs : s.Finite) : + ∃ (J : Submodule R N) (_ : J.FG) (hle : J ≤ I), + s ⊆ LinearMap.range (rTensor M (J.inclusion hle)) := by + choose J fg hle y eq using exists_fg_le_eq_rTensor_inclusion (M := M) (I := I) + rw [← Set.finite_coe_iff] at hs + refine ⟨⨆ x : s, J x, fg_iSup _ fun _ ↦ fg _, iSup_le fun _ ↦ hle _, fun x hx ↦ + ⟨rTensor M (inclusion <| le_iSup _ ⟨x, hx⟩) (y x), .trans ?_ (eq x).symm⟩⟩ + rw [← comp_apply, ← rTensor_comp]; rfl end Submodule diff --git a/Mathlib/RingTheory/Flat/CategoryTheory.lean b/Mathlib/RingTheory/Flat/CategoryTheory.lean index e8aaed01f1d4d..c29e5a8192d97 100644 --- a/Mathlib/RingTheory/Flat/CategoryTheory.lean +++ b/Mathlib/RingTheory/Flat/CategoryTheory.lean @@ -49,7 +49,7 @@ lemma rTensor_shortComplex_exact [Flat R M] (C : ShortComplex <| ModuleCat R) (h 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 <| + ⟨fun _ _ ↦ lTensor_shortComplex_exact _ _, fun H ↦ iff_lTensor_exact.2 fun _ _ _ _ _ _ _ _ _ f g h ↦ moduleCat_exact_iff_function_exact _ |>.1 <| H (.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) @@ -59,7 +59,7 @@ lemma iff_lTensor_preserves_shortComplex_exact : 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 <| + ⟨fun _ _ ↦ rTensor_shortComplex_exact _ _, fun H ↦ iff_rTensor_exact.2 fun _ _ _ _ _ _ _ _ _ f g h ↦ moduleCat_exact_iff_function_exact _ |>.1 <| H (.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) diff --git a/Mathlib/RingTheory/Localization/BaseChange.lean b/Mathlib/RingTheory/Localization/BaseChange.lean index f779f7ac70acb..2290a0b5664c5 100644 --- a/Mathlib/RingTheory/Localization/BaseChange.lean +++ b/Mathlib/RingTheory/Localization/BaseChange.lean @@ -52,13 +52,16 @@ theorem isLocalizedModule_iff_isBaseChange : IsLocalizedModule S f ↔ IsBaseCha namespace IsLocalization -include S open TensorProduct Algebra.TensorProduct +instance tensorProduct_isLocalizedModule : IsLocalizedModule S (TensorProduct.mk R A M 1) := + (isLocalizedModule_iff_isBaseChange _ A _).mpr (TensorProduct.isBaseChange _ _ _) + variable (M₁ M₂ B C) [AddCommMonoid M₁] [AddCommMonoid M₂] [Module R M₁] [Module R M₂] [Module A M₁] [Module A M₂] [IsScalarTower R A M₁] [IsScalarTower R A M₂] [Semiring B] [Algebra R B] [Algebra A B] [IsScalarTower R A B] [Semiring C] [Algebra R C] [Algebra A C] [IsScalarTower R A C] +include S theorem tensorProduct_compatibleSMul : CompatibleSMul R A M₁ M₂ where smul_tmul a _ _ := by From 8b781ec918a891323c20be4ed923133e287d624c Mon Sep 17 00:00:00 2001 From: Artie Khovanov Date: Thu, 2 Jan 2025 11:03:45 +0000 Subject: [PATCH 072/189] feat(Algebra): add missing substructure lemmas (#20269) Add missing boilerplate lemmas about the maps `.toNonUnitalSubsemiring ` and `.subsemiringClosure` This PR is split off from #16094 --- .../Algebra/Subalgebra/Unitization.lean | 30 +++++++++++++++++-- Mathlib/Algebra/Ring/Subsemiring/Basic.lean | 4 +++ 2 files changed, 32 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Unitization.lean b/Mathlib/Algebra/Algebra/Subalgebra/Unitization.lean index 2cacd22757a8c..e1bf0603035cd 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Unitization.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Unitization.lean @@ -208,10 +208,36 @@ variable {R : Type*} [NonAssocSemiring R] def Subsemiring.toNonUnitalSubsemiring (S : Subsemiring R) : NonUnitalSubsemiring R := { S with } +theorem Subsemiring.toNonUnitalSubsemiring_injective : + Function.Injective (toNonUnitalSubsemiring : Subsemiring R → _) := + fun S₁ S₂ h => SetLike.ext'_iff.2 ( + show (S₁.toNonUnitalSubsemiring : Set R) = S₂ from SetLike.ext'_iff.1 h) + +@[simp] +theorem Subsemiring.toNonUnitalSubsemiring_inj {S₁ S₂ : Subsemiring R} : + S₁.toNonUnitalSubsemiring = S₂.toNonUnitalSubsemiring ↔ S₁ = S₂ := + toNonUnitalSubsemiring_injective.eq_iff + +@[simp] +theorem Subsemiring.mem_toNonUnitalSubsemiring {S : Subsemiring R} + {x : R} : x ∈ S.toNonUnitalSubsemiring ↔ x ∈ S := Iff.rfl + +@[simp] +theorem Subsemiring.coe_toNonUnitalSubsemiring (S : Subsemiring R) : + (S.toNonUnitalSubsemiring : Set R) = S := rfl + theorem Subsemiring.one_mem_toNonUnitalSubsemiring (S : Subsemiring R) : (1 : R) ∈ S.toNonUnitalSubsemiring := S.one_mem +@[simp] +theorem Submonoid.subsemiringClosure_toNonUnitalSubsemiring {M : Submonoid R} : + M.subsemiringClosure.toNonUnitalSubsemiring = .closure M := by + refine Eq.symm (NonUnitalSubsemiring.closure_eq_of_le ?_ (fun _ hx => ?_)) + · simp [Submonoid.subsemiringClosure_coe] + · simp [Submonoid.subsemiringClosure_mem] at hx + induction hx using AddSubmonoid.closure_induction <;> aesop + /-- Turn a non-unital subsemiring containing `1` into a subsemiring. -/ def NonUnitalSubsemiring.toSubsemiring (S : NonUnitalSubsemiring R) (h1 : (1 : R) ∈ S) : Subsemiring R := @@ -241,7 +267,7 @@ theorem unitization_apply (x : Unitization ℕ s) : unitization s x = x.fst + x. rfl theorem unitization_range : - (unitization s).range = subalgebraOfSubsemiring (Subsemiring.closure s) := by + (unitization s).range = subalgebraOfSubsemiring (.closure s) := by have := AddSubmonoidClass.nsmulMemClass (S := S) rw [unitization, NonUnitalSubalgebra.unitization_range (hSRA := this), Algebra.adjoin_nat] @@ -288,7 +314,7 @@ theorem unitization_apply (x : Unitization ℤ s) : unitization s x = x.fst + x. rfl theorem unitization_range : - (unitization s).range = subalgebraOfSubring (Subring.closure s) := by + (unitization s).range = subalgebraOfSubring (.closure s) := by have := AddSubgroupClass.zsmulMemClass (S := S) rw [unitization, NonUnitalSubalgebra.unitization_range (hSRA := this), Algebra.adjoin_int] diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index 88035633f260c..db80142d9f7af 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -391,6 +391,10 @@ theorem subsemiringClosure_coe : (M.subsemiringClosure : Set R) = AddSubmonoid.closure (M : Set R) := rfl +theorem subsemiringClosure_mem {x : R} : + x ∈ M.subsemiringClosure ↔ x ∈ AddSubmonoid.closure (M : Set R) := + Iff.rfl + theorem subsemiringClosure_toAddSubmonoid : M.subsemiringClosure.toAddSubmonoid = AddSubmonoid.closure (M : Set R) := rfl From 5f024868c0081c59a83baa4f4d6c84c7431e8697 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 2 Jan 2025 11:03:46 +0000 Subject: [PATCH 073/189] feat(Order/Filter): add `Ultrafilter.eventually_exists_iff` (#20280) --- Mathlib/Order/Filter/Ultrafilter.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Mathlib/Order/Filter/Ultrafilter.lean b/Mathlib/Order/Filter/Ultrafilter.lean index e480d43fb294c..d171f958a9ca2 100644 --- a/Mathlib/Order/Filter/Ultrafilter.lean +++ b/Mathlib/Order/Filter/Ultrafilter.lean @@ -168,6 +168,16 @@ theorem finite_biUnion_mem_iff {is : Set β} {s : β → Set α} (his : is.Finit (⋃ i ∈ is, s i) ∈ f ↔ ∃ i ∈ is, s i ∈ f := by simp only [← sUnion_image, finite_sUnion_mem_iff (his.image s), exists_mem_image] +lemma eventually_exists_mem_iff {is : Set β} {P : β → α → Prop} (his : is.Finite) : + (∀ᶠ i in f, ∃ a ∈ is, P a i) ↔ ∃ a ∈ is, ∀ᶠ i in f, P a i := by + simp only [Filter.Eventually, Ultrafilter.mem_coe] + convert f.finite_biUnion_mem_iff his (s := P) with i + aesop + +lemma eventually_exists_iff [Finite β] {P : β → α → Prop} : + (∀ᶠ i in f, ∃ a, P a i) ↔ ∃ a, ∀ᶠ i in f, P a i := by + simpa using eventually_exists_mem_iff (f := f) (P := P) Set.finite_univ + /-- Pushforward for ultrafilters. -/ nonrec def map (m : α → β) (f : Ultrafilter α) : Ultrafilter β := ofComplNotMemIff (map m f) fun s => @compl_not_mem_iff _ f (m ⁻¹' s) From 952a09f9423a89e64be64ab58e3f956ed8328bbe Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 2 Jan 2025 11:03:47 +0000 Subject: [PATCH 074/189] feat(Order/Filter): add `Filter.skolem` (#20283) --- Mathlib/Order/Filter/Basic.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index 7314f90b81bb6..6dd554714e2b1 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -848,6 +848,15 @@ theorem Eventually.choice {r : α → β → Prop} {l : Filter α} [l.NeBot] (h choose! f hf using fun x (hx : ∃ y, r x y) => hx exact ⟨f, h.mono hf⟩ +lemma skolem {ι : Type*} {α : ι → Type*} [∀ i, Nonempty (α i)] + {P : ∀ i : ι, α i → Prop} {F : Filter ι} : + (∀ᶠ i in F, ∃ b, P i b) ↔ ∃ b : (Π i, α i), ∀ᶠ i in F, P i (b i) := by + classical + refine ⟨fun H ↦ ?_, fun ⟨b, hb⟩ ↦ hb.mp (.of_forall fun x a ↦ ⟨_, a⟩)⟩ + refine ⟨fun i ↦ if h : ∃ b, P i b then h.choose else Nonempty.some inferInstance, ?_⟩ + filter_upwards [H] with i hi + exact dif_pos hi ▸ hi.choose_spec + /-! ### Relation “eventually equal” -/ From 143b4dd8da635e7922fe4f95fc003428713f62f0 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 2 Jan 2025 11:03:48 +0000 Subject: [PATCH 075/189] chore: remove the superflous `NormedSpace.toModule'` and `NormedAlgebra.toNormedSpace'` (#20349) This was [relevant in Lean 3](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/Typeclass.20resolution.20under.20binders/near/245151099), but everything just works now. --- Mathlib/Analysis/Analytic/OfScalars.lean | 3 ++ Mathlib/Analysis/Calculus/ContDiff/Basic.lean | 10 ++--- Mathlib/Analysis/Calculus/ContDiff/Defs.lean | 2 +- .../Calculus/ContDiff/FTaylorSeries.lean | 2 +- Mathlib/Analysis/InnerProductSpace/Basic.lean | 4 +- Mathlib/Analysis/Normed/Module/Basic.lean | 42 ------------------- 6 files changed, 12 insertions(+), 51 deletions(-) diff --git a/Mathlib/Analysis/Analytic/OfScalars.lean b/Mathlib/Analysis/Analytic/OfScalars.lean index dfd0330aa12e1..771e16242e8ad 100644 --- a/Mathlib/Analysis/Analytic/OfScalars.lean +++ b/Mathlib/Analysis/Analytic/OfScalars.lean @@ -147,6 +147,9 @@ open scoped Topology NNReal variable {𝕜 : Type*} (E : Type*) [NontriviallyNormedField 𝕜] [NormedRing E] [NormedAlgebra 𝕜 E] (c : ℕ → 𝕜) (n : ℕ) +-- Also works: +-- `letI : BoundedSMul 𝕜 (ContinuousMultilinearMap 𝕜 (fun i : Fin n ↦ E) E) := inferInstance` +set_option maxSynthPendingDepth 2 in theorem ofScalars_norm_eq_mul : ‖ofScalars E c n‖ = ‖c n‖ * ‖ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 n E‖ := by rw [ofScalars, norm_smul (c n) (ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 n E)] diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 2cab2349803e7..87efceee82a5c 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -40,7 +40,7 @@ open scoped NNReal Nat ContDiff universe u uE uF uG attribute [local instance 1001] - NormedAddCommGroup.toAddCommGroup NormedSpace.toModule' AddCommGroup.toAddCommMonoid + NormedAddCommGroup.toAddCommGroup AddCommGroup.toAddCommMonoid open Set Fin Filter Function @@ -2046,10 +2046,10 @@ theorem HasFTaylorSeriesUpToOn.restrictScalars {n : WithTop ℕ∞} (h : HasFTaylorSeriesUpToOn n f p' s) : HasFTaylorSeriesUpToOn n f (fun x => (p' x).restrictScalars 𝕜) s where zero_eq x hx := h.zero_eq x hx - fderivWithin m hm x hx := by - simpa only using -- Porting note: added `by simpa only using` - (ContinuousMultilinearMap.restrictScalarsLinear 𝕜).hasFDerivAt.comp_hasFDerivWithinAt x <| - (h.fderivWithin m hm x hx).restrictScalars 𝕜 + fderivWithin m hm x hx := + set_option maxSynthPendingDepth 2 in + ((ContinuousMultilinearMap.restrictScalarsLinear 𝕜).hasFDerivAt.comp_hasFDerivWithinAt x <| + (h.fderivWithin m hm x hx).restrictScalars 𝕜 :) cont m hm := ContinuousMultilinearMap.continuous_restrictScalars.comp_continuousOn (h.cont m hm) theorem ContDiffWithinAt.restrict_scalars (h : ContDiffWithinAt 𝕜' n f s x) : diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index 6fbd690d334d6..eabcdff56fab9 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -101,7 +101,7 @@ open NNReal Topology Filter /- Porting note: These lines are not required in Mathlib4. attribute [local instance 1001] - NormedAddCommGroup.toAddCommGroup NormedSpace.toModule' AddCommGroup.toAddCommMonoid + NormedAddCommGroup.toAddCommGroup AddCommGroup.toAddCommMonoid -/ open Set Fin Filter Function diff --git a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean index ea54387d8a7cd..400df9cb52915 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean @@ -105,7 +105,7 @@ open ENat NNReal Topology Filter Set Fin Filter Function /- Porting note: These lines are not required in Mathlib4. attribute [local instance 1001] - NormedAddCommGroup.toAddCommGroup NormedSpace.toModule' AddCommGroup.toAddCommMonoid + NormedAddCommGroup.toAddCommGroup AddCommGroup.toAddCommMonoid -/ /-- Smoothness exponent for analytic functions. -/ diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 3bc54c60335ff..4ec4b09eb713c 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -703,8 +703,8 @@ theorem norm_inner_eq_norm_tfae (x y : E) : try positivity simp only [@norm_sq_eq_inner 𝕜] at h letI : InnerProductSpace.Core 𝕜 E := InnerProductSpace.toCore - erw [← InnerProductSpace.Core.cauchy_schwarz_aux, InnerProductSpace.Core.normSq_eq_zero, - sub_eq_zero] at h + erw [← InnerProductSpace.Core.cauchy_schwarz_aux (𝕜 := 𝕜) (F := E), + InnerProductSpace.Core.normSq_eq_zero, sub_eq_zero] at h rw [div_eq_inv_mul, mul_smul, h, inv_smul_smul₀] rwa [inner_self_ne_zero] tfae_have 2 → 3 := fun h => h.imp_right fun h' => ⟨_, h'⟩ diff --git a/Mathlib/Analysis/Normed/Module/Basic.lean b/Mathlib/Analysis/Normed/Module/Basic.lean index 717db36576e71..615129ab6a77a 100644 --- a/Mathlib/Analysis/Normed/Module/Basic.lean +++ b/Mathlib/Analysis/Normed/Module/Basic.lean @@ -150,33 +150,6 @@ abbrev NormedSpace.induced {F : Type*} (𝕜 E G : Type*) [NormedField 𝕜] [Ad let _ := SeminormedAddCommGroup.induced E G f ⟨fun a b ↦ by simpa only [← map_smul f a b] using norm_smul_le a (f b)⟩ -section NormedAddCommGroup - -variable [NormedField 𝕜] -variable [NormedAddCommGroup E] [NormedSpace 𝕜 E] -variable [NormedAddCommGroup F] [NormedSpace 𝕜 F] - -open NormedField - -/-- While this may appear identical to `NormedSpace.toModule`, it contains an implicit argument -involving `NormedAddCommGroup.toSeminormedAddCommGroup` that typeclass inference has trouble -inferring. - -Specifically, the following instance cannot be found without this `NormedSpace.toModule'`: -```lean -example - (𝕜 ι : Type*) (E : ι → Type*) - [NormedField 𝕜] [Π i, NormedAddCommGroup (E i)] [Π i, NormedSpace 𝕜 (E i)] : - Π i, Module 𝕜 (E i) := by infer_instance -``` - -[This Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/Typeclass.20resolution.20under.20binders/near/245151099) -gives some more context. -/ -instance (priority := 100) NormedSpace.toModule' : Module 𝕜 F := - NormedSpace.toModule - -end NormedAddCommGroup - section NontriviallyNormedSpace variable (𝕜 E) @@ -267,21 +240,6 @@ instance (priority := 100) NormedAlgebra.toNormedSpace : NormedSpace 𝕜 𝕜' { NormedAlgebra.toAlgebra.toModule with norm_smul_le := NormedAlgebra.norm_smul_le } -/-- While this may appear identical to `NormedAlgebra.toNormedSpace`, it contains an implicit -argument involving `NormedRing.toSeminormedRing` that typeclass inference has trouble inferring. - -Specifically, the following instance cannot be found without this `NormedSpace.toModule'`: -```lean -example - (𝕜 ι : Type*) (E : ι → Type*) - [NormedField 𝕜] [Π i, NormedRing (E i)] [Π i, NormedAlgebra 𝕜 (E i)] : - Π i, Module 𝕜 (E i) := by infer_instance -``` - -See `NormedSpace.toModule'` for a similar situation. -/ -instance (priority := 100) NormedAlgebra.toNormedSpace' {𝕜'} [NormedRing 𝕜'] [NormedAlgebra 𝕜 𝕜'] : - NormedSpace 𝕜 𝕜' := by infer_instance - theorem norm_algebraMap (x : 𝕜) : ‖algebraMap 𝕜 𝕜' x‖ = ‖x‖ * ‖(1 : 𝕜')‖ := by rw [Algebra.algebraMap_eq_smul_one] exact norm_smul _ _ From a07a33c81892c3b510ff9980e7fe7b8f8623006e Mon Sep 17 00:00:00 2001 From: Artie Khovanov Date: Thu, 2 Jan 2025 11:03:50 +0000 Subject: [PATCH 076/189] chore(Mathlib/Algebra/Lie): rename theorems for consistency (#20353) Moves: - `*coeSubmodule*`, `*coe_toSubmodule*`, `*to_submodule*`, `*coe_submodule*`, `*coe_to_submodule*` -> `*toSubmodule*` - `*coe_linearMap*` -> `*toLinearMap*` - `*to_lieHom*` -> `*toLieHom*` - `*coe_linearEquiv*`, `*to_linearEquiv*` -> `*toLinearEquiv*` - `*coe_lieSubalgebra*`, `*to_lieSubalgebra*` -> `*toLieSubalgebra*` special cases: * `LieIdeal.coe_toSubalgebra` -> `LieIdeal.coe_toLieSubalgebra` * `LieSubalgebra.coe_to_submodule` -> `LieSubalgebra.coe_toSubmodule` * `LieSubmodule.coe_toSubmodule` and similar: unchanged * `sInf_coe_toSubmodule'` -> `sInf_toSubmodule_eq_iInf` * `sSup_coe_toSubmodule'` -> `sSup_toSubmodule_eq_iSup` PR follows [this Zulip discussion](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Algebra.2ELie.2ESubmodule.2C.20Algebra.2ELie.2ESubalgebra.20naming) --- Mathlib/Algebra/Lie/Abelian.lean | 20 +- Mathlib/Algebra/Lie/BaseChange.lean | 6 +- Mathlib/Algebra/Lie/Basic.lean | 36 ++- Mathlib/Algebra/Lie/CartanExists.lean | 9 +- Mathlib/Algebra/Lie/CartanSubalgebra.lean | 6 +- Mathlib/Algebra/Lie/Character.lean | 2 +- Mathlib/Algebra/Lie/Derivation/Killing.lean | 4 +- Mathlib/Algebra/Lie/Engel.lean | 20 +- Mathlib/Algebra/Lie/EngelSubalgebra.lean | 4 +- Mathlib/Algebra/Lie/IdealOperations.lean | 10 +- Mathlib/Algebra/Lie/InvariantForm.lean | 19 +- Mathlib/Algebra/Lie/Nilpotent.lean | 28 +- Mathlib/Algebra/Lie/Normalizer.lean | 4 +- Mathlib/Algebra/Lie/OfAssociative.lean | 4 +- Mathlib/Algebra/Lie/Semisimple/Basic.lean | 6 +- Mathlib/Algebra/Lie/Subalgebra.lean | 55 ++-- Mathlib/Algebra/Lie/Submodule.lean | 267 ++++++++++++------- Mathlib/Algebra/Lie/TensorProduct.lean | 14 +- Mathlib/Algebra/Lie/TraceForm.lean | 16 +- Mathlib/Algebra/Lie/UniversalEnveloping.lean | 2 +- Mathlib/Algebra/Lie/Weights/Basic.lean | 26 +- Mathlib/Algebra/Lie/Weights/Cartan.lean | 12 +- Mathlib/Algebra/Lie/Weights/Chain.lean | 4 +- Mathlib/Algebra/Lie/Weights/Killing.lean | 18 +- Mathlib/Algebra/Lie/Weights/Linear.lean | 8 +- 25 files changed, 363 insertions(+), 237 deletions(-) diff --git a/Mathlib/Algebra/Lie/Abelian.lean b/Mathlib/Algebra/Lie/Abelian.lean index 3f1dcaf33206a..78f6cb9140859 100644 --- a/Mathlib/Algebra/Lie/Abelian.lean +++ b/Mathlib/Algebra/Lie/Abelian.lean @@ -121,8 +121,8 @@ instance : IsTrivial L (maxTrivSubmodule R L M) where trivial x m := Subtype.ext @[simp] theorem ideal_oper_maxTrivSubmodule_eq_bot (I : LieIdeal R L) : ⁅I, maxTrivSubmodule R L M⁆ = ⊥ := by - rw [← LieSubmodule.coe_toSubmodule_inj, LieSubmodule.lieIdeal_oper_eq_linear_span, - LieSubmodule.bot_coeSubmodule, Submodule.span_eq_bot] + rw [← LieSubmodule.toSubmodule_inj, LieSubmodule.lieIdeal_oper_eq_linear_span, + LieSubmodule.bot_toSubmodule, Submodule.span_eq_bot] rintro m ⟨⟨x, hx⟩, ⟨⟨m, hm⟩, rfl⟩⟩ exact hm x @@ -167,8 +167,8 @@ def maxTrivEquiv (e : M ≃ₗ⁅R,L⁆ N) : maxTrivSubmodule R L M ≃ₗ⁅R,L { maxTrivHom (e : M →ₗ⁅R,L⁆ N) with toFun := maxTrivHom (e : M →ₗ⁅R,L⁆ N) invFun := maxTrivHom (e.symm : N →ₗ⁅R,L⁆ M) - left_inv := fun m => by ext; simp [LieModuleEquiv.coe_to_lieModuleHom] - right_inv := fun n => by ext; simp [LieModuleEquiv.coe_to_lieModuleHom] } + left_inv := fun m => by ext; simp [LieModuleEquiv.coe_toLieModuleHom] + right_inv := fun n => by ext; simp [LieModuleEquiv.coe_toLieModuleHom] } @[norm_cast, simp] theorem coe_maxTrivEquiv_apply (e : M ≃ₗ⁅R,L⁆ N) (m : maxTrivSubmodule R L M) : @@ -209,15 +209,23 @@ theorem coe_maxTrivLinearMapEquivLieModuleHom_symm (f : M →ₗ⁅R,L⁆ N) : rfl @[simp] -theorem coe_linearMap_maxTrivLinearMapEquivLieModuleHom (f : maxTrivSubmodule R L (M →ₗ[R] N)) : +theorem toLinearMap_maxTrivLinearMapEquivLieModuleHom (f : maxTrivSubmodule R L (M →ₗ[R] N)) : (maxTrivLinearMapEquivLieModuleHom (M := M) (N := N) f : M →ₗ[R] N) = (f : M →ₗ[R] N) := by ext; rfl +@[deprecated (since := "2024-12-30")] +alias coe_linearMap_maxTrivLinearMapEquivLieModuleHom := + toLinearMap_maxTrivLinearMapEquivLieModuleHom + @[simp] -theorem coe_linearMap_maxTrivLinearMapEquivLieModuleHom_symm (f : M →ₗ⁅R,L⁆ N) : +theorem toLinearMap_maxTrivLinearMapEquivLieModuleHom_symm (f : M →ₗ⁅R,L⁆ N) : (maxTrivLinearMapEquivLieModuleHom (M := M) (N := N) |>.symm f : M →ₗ[R] N) = (f : M →ₗ[R] N) := rfl +@[deprecated (since := "2024-12-30")] +alias coe_linearMap_maxTrivLinearMapEquivLieModuleHom_symm := + toLinearMap_maxTrivLinearMapEquivLieModuleHom_symm + end LieModule namespace LieAlgebra diff --git a/Mathlib/Algebra/Lie/BaseChange.lean b/Mathlib/Algebra/Lie/BaseChange.lean index 5a8ada230004c..d6be0876fe3d9 100644 --- a/Mathlib/Algebra/Lie/BaseChange.lean +++ b/Mathlib/Algebra/Lie/BaseChange.lean @@ -200,13 +200,13 @@ lemma mem_baseChange_iff {m : A ⊗[R] M} : @[simp] lemma baseChange_bot : (⊥ : LieSubmodule R L M).baseChange A = ⊥ := by - simp only [baseChange, bot_coeSubmodule, Submodule.baseChange_bot, + simp only [baseChange, bot_toSubmodule, Submodule.baseChange_bot, Submodule.bot_toAddSubmonoid] rfl @[simp] lemma baseChange_top : (⊤ : LieSubmodule R L M).baseChange A = ⊤ := by - simp only [baseChange, top_coeSubmodule, Submodule.baseChange_top, + simp only [baseChange, top_toSubmodule, Submodule.baseChange_top, Submodule.bot_toAddSubmonoid] rfl @@ -214,7 +214,7 @@ lemma lie_baseChange {I : LieIdeal R L} {N : LieSubmodule R L M} : ⁅I, N⁆.baseChange A = ⁅I.baseChange A, N.baseChange A⁆ := by set s : Set (A ⊗[R] M) := { m | ∃ x ∈ I, ∃ n ∈ N, 1 ⊗ₜ ⁅x, n⁆ = m} have : (TensorProduct.mk R A M 1) '' {m | ∃ x ∈ I, ∃ n ∈ N, ⁅x, n⁆ = m} = s := by ext; simp [s] - rw [← coe_toSubmodule_inj, coe_baseChange, lieIdeal_oper_eq_linear_span', + rw [← toSubmodule_inj, coe_baseChange, lieIdeal_oper_eq_linear_span', Submodule.baseChange_span, this, lieIdeal_oper_eq_linear_span'] refine le_antisymm (Submodule.span_mono ?_) (Submodule.span_le.mpr ?_) · rintro - ⟨x, hx, m, hm, rfl⟩ diff --git a/Mathlib/Algebra/Lie/Basic.lean b/Mathlib/Algebra/Lie/Basic.lean index cb37eac2841ca..8ff00bfb403c6 100644 --- a/Mathlib/Algebra/Lie/Basic.lean +++ b/Mathlib/Algebra/Lie/Basic.lean @@ -380,10 +380,12 @@ theorem coe_comp (f : L₂ →ₗ⁅R⁆ L₃) (g : L₁ →ₗ⁅R⁆ L₂) : ( rfl @[norm_cast, simp] -theorem coe_linearMap_comp (f : L₂ →ₗ⁅R⁆ L₃) (g : L₁ →ₗ⁅R⁆ L₂) : +theorem toLinearMap_comp (f : L₂ →ₗ⁅R⁆ L₃) (g : L₁ →ₗ⁅R⁆ L₂) : (f.comp g : L₁ →ₗ[R] L₃) = (f : L₂ →ₗ[R] L₃).comp (g : L₁ →ₗ[R] L₂) := rfl +@[deprecated (since := "2024-12-30")] alias coe_linearMap_comp := toLinearMap_comp + @[simp] theorem comp_id (f : L₁ →ₗ⁅R⁆ L₂) : f.comp (id : L₁ →ₗ⁅R⁆ L₁) = f := rfl @@ -478,15 +480,19 @@ instance : EquivLike (L₁ ≃ₗ⁅R⁆ L₂) L₁ L₂ where right_inv f := f.right_inv coe_injective' f g h₁ h₂ := by cases f; cases g; simp at h₁ h₂; simp [*] -theorem coe_to_lieHom (e : L₁ ≃ₗ⁅R⁆ L₂) : ⇑(e : L₁ →ₗ⁅R⁆ L₂) = e := +theorem coe_toLieHom (e : L₁ ≃ₗ⁅R⁆ L₂) : ⇑(e : L₁ →ₗ⁅R⁆ L₂) = e := rfl +@[deprecated (since := "2024-12-30")] alias coe_to_lieHom := coe_toLieHom + @[simp] -theorem coe_to_linearEquiv (e : L₁ ≃ₗ⁅R⁆ L₂) : ⇑(e : L₁ ≃ₗ[R] L₂) = e := +theorem coe_toLinearEquiv (e : L₁ ≃ₗ⁅R⁆ L₂) : ⇑(e : L₁ ≃ₗ[R] L₂) = e := rfl +@[deprecated (since := "2024-12-30")] alias coe_to_linearEquiv := coe_toLinearEquiv + @[simp] -theorem to_linearEquiv_mk (f : L₁ →ₗ⁅R⁆ L₂) (g h₁ h₂) : +theorem toLinearEquiv_mk (f : L₁ →ₗ⁅R⁆ L₂) (g h₁ h₂) : (mk f g h₁ h₂ : L₁ ≃ₗ[R] L₂) = { f with invFun := g @@ -494,15 +500,19 @@ theorem to_linearEquiv_mk (f : L₁ →ₗ⁅R⁆ L₂) (g h₁ h₂) : right_inv := h₂ } := rfl -theorem coe_linearEquiv_injective : Injective ((↑) : (L₁ ≃ₗ⁅R⁆ L₂) → L₁ ≃ₗ[R] L₂) := by +@[deprecated (since := "2024-12-30")] alias to_linearEquiv_mk := toLinearEquiv_mk + +theorem toLinearEquiv_injective : Injective ((↑) : (L₁ ≃ₗ⁅R⁆ L₂) → L₁ ≃ₗ[R] L₂) := by rintro ⟨⟨⟨⟨f, -⟩, -⟩, -⟩, f_inv⟩ ⟨⟨⟨⟨g, -⟩, -⟩, -⟩, g_inv⟩ intro h - simp only [to_linearEquiv_mk, LinearEquiv.mk.injEq, LinearMap.mk.injEq, AddHom.mk.injEq] at h + simp only [toLinearEquiv_mk, LinearEquiv.mk.injEq, LinearMap.mk.injEq, AddHom.mk.injEq] at h congr exacts [h.1, h.2] +@[deprecated (since := "2024-12-30")] alias coe_linearEquiv_injective := toLinearEquiv_injective + theorem coe_injective : @Injective (L₁ ≃ₗ⁅R⁆ L₂) (L₁ → L₂) (↑) := - LinearEquiv.coe_injective.comp coe_linearEquiv_injective + LinearEquiv.coe_injective.comp toLinearEquiv_injective @[ext] theorem ext {f g : L₁ ≃ₗ⁅R⁆ L₂} (h : ∀ x, f x = g x) : f = g := @@ -727,10 +737,12 @@ theorem coe_comp (f : N →ₗ⁅R,L⁆ P) (g : M →ₗ⁅R,L⁆ N) : ⇑(f.com rfl @[norm_cast, simp] -theorem coe_linearMap_comp (f : N →ₗ⁅R,L⁆ P) (g : M →ₗ⁅R,L⁆ N) : +theorem toLinearMap_comp (f : N →ₗ⁅R,L⁆ P) (g : M →ₗ⁅R,L⁆ N) : (f.comp g : M →ₗ[R] P) = (f : N →ₗ[R] P).comp (g : M →ₗ[R] N) := rfl +@[deprecated (since := "2024-12-30")] alias coe_linearMap_comp := toLinearMap_comp + /-- The inverse of a bijective morphism of Lie modules is a morphism of Lie modules. -/ def inverse (f : M →ₗ⁅R,L⁆ N) (g : N → M) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) : N →ₗ⁅R,L⁆ M := @@ -878,13 +890,17 @@ theorem coe_mk (f : M →ₗ⁅R,L⁆ N) (invFun h₁ h₂) : ((⟨f, invFun, h₁, h₂⟩ : M ≃ₗ⁅R,L⁆ N) : M → N) = f := rfl -theorem coe_to_lieModuleHom (e : M ≃ₗ⁅R,L⁆ N) : ⇑(e : M →ₗ⁅R,L⁆ N) = e := +theorem coe_toLieModuleHom (e : M ≃ₗ⁅R,L⁆ N) : ⇑(e : M →ₗ⁅R,L⁆ N) = e := rfl +@[deprecated (since := "2024-12-30")] alias coe_to_lieModuleHom := coe_toLieModuleHom + @[simp] -theorem coe_to_linearEquiv (e : M ≃ₗ⁅R,L⁆ N) : ((e : M ≃ₗ[R] N) : M → N) = e := +theorem coe_toLinearEquiv (e : M ≃ₗ⁅R,L⁆ N) : ((e : M ≃ₗ[R] N) : M → N) = e := rfl +@[deprecated (since := "2024-12-30")] alias coe_to_linearEquiv := coe_toLinearEquiv + theorem toEquiv_injective : Function.Injective (toEquiv : (M ≃ₗ⁅R,L⁆ N) → M ≃ N) := by rintro ⟨⟨⟨⟨f, -⟩, -⟩, -⟩, f_inv⟩ ⟨⟨⟨⟨g, -⟩, -⟩, -⟩, g_inv⟩ intro h diff --git a/Mathlib/Algebra/Lie/CartanExists.lean b/Mathlib/Algebra/Lie/CartanExists.lean index cca8926c878a6..4e3acbb133d81 100644 --- a/Mathlib/Algebra/Lie/CartanExists.lean +++ b/Mathlib/Algebra/Lie/CartanExists.lean @@ -152,7 +152,7 @@ lemma engel_isBot_of_isMin (hLK : finrank K L ≤ #K) (U : LieSubalgebra K L) obtain hr|hr : r = finrank K L ∨ r < finrank K L := (Submodule.finrank_le _).eq_or_lt · suffices engel K y ≤ engel K x from hmin Ey this suffices engel K x = ⊤ by simp_rw [this, le_top] - apply LieSubalgebra.to_submodule_injective + apply LieSubalgebra.toSubmodule_injective apply Submodule.eq_top_of_finrank_eq hr -- So from now on, we assume that `r < finrank K L`. -- We denote by `x'` and `y'` the elements `x` and `y` viewed as terms of `U`. @@ -246,8 +246,9 @@ lemma engel_isBot_of_isMin (hLK : finrank K L ≤ #K) (U : LieSubalgebra K L) -- From this we deduce that there exists an `n` such that `⁅x, _⁆ ^ n` vanishes on `⁅x, z⁆`. -- On the other hand, our goal is to show `z = 0` in `Q`, -- which is equivalent to showing that `⁅x, _⁆ ^ n` vanishes on `z`, for some `n`. - simp only [coe_bracket_of_module, LieSubmodule.mem_mk_iff', mem_coe_submodule, mem_engel_iff, - LieSubmodule.Quotient.mk'_apply, LieSubmodule.Quotient.mk_eq_zero', E, Q] at this ⊢ + simp only [coe_bracket_of_module, LieSubmodule.mem_mk_iff', LieSubalgebra.mem_toSubmodule, + mem_engel_iff, LieSubmodule.Quotient.mk'_apply, LieSubmodule.Quotient.mk_eq_zero', E, Q] + at this ⊢ -- Hence we win. obtain ⟨n, hn⟩ := this use n+1 @@ -359,7 +360,7 @@ lemma exists_isCartanSubalgebra_engel_of_finrank_le_card (h : finrank K L ≤ #K rintro ⟨_, y, hy, rfl⟩ hyx suffices finrank K (engel K x) ≤ finrank K (engel K y) by suffices engel K y = engel K x from this.ge - apply LieSubalgebra.to_submodule_injective + apply LieSubalgebra.toSubmodule_injective exact Submodule.eq_of_le_of_finrank_le hyx this rw [(isRegular_iff_finrank_engel_eq_rank K x).mp hx] apply rank_le_finrank_engel diff --git a/Mathlib/Algebra/Lie/CartanSubalgebra.lean b/Mathlib/Algebra/Lie/CartanSubalgebra.lean index fed2b9e12669c..cfe41bd7c8bfd 100644 --- a/Mathlib/Algebra/Lie/CartanSubalgebra.lean +++ b/Mathlib/Algebra/Lie/CartanSubalgebra.lean @@ -53,7 +53,7 @@ instance [H.IsCartanSubalgebra] : LieAlgebra.IsNilpotent R H := @[simp] theorem normalizer_eq_self_of_isCartanSubalgebra (H : LieSubalgebra R L) [H.IsCartanSubalgebra] : H.toLieSubmodule.normalizer = H.toLieSubmodule := by - rw [← LieSubmodule.coe_toSubmodule_inj, coe_normalizer_eq_normalizer, + rw [← LieSubmodule.toSubmodule_inj, coe_normalizer_eq_normalizer, IsCartanSubalgebra.self_normalizing, coe_toLieSubmodule] @[simp] @@ -84,7 +84,7 @@ theorem isCartanSubalgebra_iff_isUcsLimit : H.IsCartanSubalgebra ↔ H.toLieSubm self_normalizing := by have hk' := hk (k + 1) k.le_succ rw [LieSubmodule.ucs_succ, hk k (le_refl k)] at hk' - rw [← LieSubalgebra.coe_to_submodule_inj, ← LieSubalgebra.coe_normalizer_eq_normalizer, + rw [← LieSubalgebra.toSubmodule_inj, ← LieSubalgebra.coe_normalizer_eq_normalizer, hk', LieSubalgebra.coe_toLieSubmodule] } lemma ne_bot_of_isCartanSubalgebra [Nontrivial L] (H : LieSubalgebra R L) [H.IsCartanSubalgebra] : @@ -116,4 +116,4 @@ open LieIdeal instance LieAlgebra.top_isCartanSubalgebra_of_nilpotent [LieAlgebra.IsNilpotent R L] : LieSubalgebra.IsCartanSubalgebra (⊤ : LieSubalgebra R L) where nilpotent := inferInstance - self_normalizing := by rw [← top_coe_lieSubalgebra, normalizer_eq_top, top_coe_lieSubalgebra] + self_normalizing := by rw [← top_toLieSubalgebra, normalizer_eq_top, top_toLieSubalgebra] diff --git a/Mathlib/Algebra/Lie/Character.lean b/Mathlib/Algebra/Lie/Character.lean index 24978ce025055..d4933f1942279 100644 --- a/Mathlib/Algebra/Lie/Character.lean +++ b/Mathlib/Algebra/Lie/Character.lean @@ -48,7 +48,7 @@ theorem lieCharacter_apply_lie' (χ : LieCharacter R L) (x y : L) : ⁅χ x, χ theorem lieCharacter_apply_of_mem_derived (χ : LieCharacter R L) {x : L} (h : x ∈ derivedSeries R L 1) : χ x = 0 := by rw [derivedSeries_def, derivedSeriesOfIdeal_succ, derivedSeriesOfIdeal_zero, ← - LieSubmodule.mem_coeSubmodule, LieSubmodule.lieIdeal_oper_eq_linear_span] at h + LieSubmodule.mem_toSubmodule, LieSubmodule.lieIdeal_oper_eq_linear_span] at h refine Submodule.span_induction ?_ ?_ ?_ ?_ h · rintro y ⟨⟨z, hz⟩, ⟨⟨w, hw⟩, rfl⟩⟩; apply lieCharacter_apply_lie · exact χ.map_zero diff --git a/Mathlib/Algebra/Lie/Derivation/Killing.lean b/Mathlib/Algebra/Lie/Derivation/Killing.lean index 04b2110b0fe93..369ad9011ee67 100644 --- a/Mathlib/Algebra/Lie/Derivation/Killing.lean +++ b/Mathlib/Algebra/Lie/Derivation/Killing.lean @@ -61,7 +61,7 @@ variable {R L} any `x : L`, `ad (D x)` is also in this orthogonal. -/ lemma ad_mem_orthogonal_of_mem_orthogonal {D : LieDerivation R L L} (hD : D ∈ 𝕀ᗮ) (x : L) : ad R L (D x) ∈ 𝕀ᗮ := by - simp only [ad_apply_lieDerivation, LieHom.range_coeSubmodule, neg_mem_iff] + simp only [ad_apply_lieDerivation, LieHom.range_toSubmodule, neg_mem_iff] exact (rangeAdOrthogonal R L).lie_mem hD variable [Module.Finite R L] @@ -94,7 +94,7 @@ lemma killingForm_restrict_range_ad_nondegenerate : /-- The range of the adjoint action on a finite-dimensional Killing Lie algebra is full. -/ @[simp] lemma range_ad_eq_top : 𝕀 = ⊤ := by - rw [← LieSubalgebra.coe_to_submodule_inj] + rw [← LieSubalgebra.toSubmodule_inj] apply LinearMap.BilinForm.eq_top_of_restrict_nondegenerate_of_orthogonal_eq_bot (LieModule.traceForm_isSymm R 𝔻 𝔻).isRefl (killingForm_restrict_range_ad_nondegenerate R L) refine (Submodule.eq_bot_iff _).mpr fun D hD ↦ ext fun x ↦ ?_ diff --git a/Mathlib/Algebra/Lie/Engel.lean b/Mathlib/Algebra/Lie/Engel.lean index 5bf6908043e53..6cab534fb7f10 100644 --- a/Mathlib/Algebra/Lie/Engel.lean +++ b/Mathlib/Algebra/Lie/Engel.lean @@ -128,7 +128,7 @@ theorem isNilpotentOfIsNilpotentSpanSupEqTop (hnp : IsNilpotent <| toEnd R L M x obtain ⟨n, hn⟩ := hnp obtain ⟨k, hk⟩ := hIM have hk' : I.lcs M k = ⊥ := by - simp only [← coe_toSubmodule_inj, I.coe_lcs_eq, hk, bot_coeSubmodule] + simp only [← toSubmodule_inj, I.coe_lcs_eq, hk, bot_toSubmodule] suffices ∀ l, lowerCentralSeries R L M (l * n) ≤ I.lcs M l by use k * n simpa [hk'] using this k @@ -189,16 +189,16 @@ theorem LieAlgebra.exists_engelian_lieSubalgebra_of_lt_normalizer {K : LieSubalg { (R ∙ x) ⊔ (K : Submodule R L) with lie_mem' := fun {y z} => LieSubalgebra.lie_mem_sup_of_mem_normalizer hx₁ } have hxK' : x ∈ K' := Submodule.mem_sup_left (Submodule.subset_span (Set.mem_singleton _)) - have hKK' : K ≤ K' := (LieSubalgebra.coe_submodule_le_coe_submodule K K').mp le_sup_right + have hKK' : K ≤ K' := (LieSubalgebra.toSubmodule_le_toSubmodule K K').mp le_sup_right have hK' : K' ≤ K.normalizer := by - rw [← LieSubalgebra.coe_submodule_le_coe_submodule] + rw [← LieSubalgebra.toSubmodule_le_toSubmodule] exact sup_le ((Submodule.span_singleton_le_iff_mem _ _).mpr hx₁) hK₂.le refine ⟨K', ?_, lt_iff_le_and_ne.mpr ⟨hKK', fun contra => hx₂ (contra.symm ▸ hxK')⟩⟩ intro M _i1 _i2 _i3 _i4 h obtain ⟨I, hI₁ : (I : LieSubalgebra R K') = LieSubalgebra.ofLe hKK'⟩ := LieSubalgebra.exists_nested_lieIdeal_ofLe_normalizer hKK' hK' have hI₂ : (R ∙ (⟨x, hxK'⟩ : K')) ⊔ (LieSubmodule.toSubmodule I) = ⊤ := by - rw [← LieIdeal.coe_to_lieSubalgebra_to_submodule R K' I, hI₁] + rw [← LieIdeal.coe_toLieSubalgebra_toSubmodule R K' I, hI₁] apply Submodule.map_injective_of_injective (K' : Submodule R L).injective_subtype simp only [LieSubalgebra.coe_ofLe, Submodule.map_sup, Submodule.map_subtype_range_inclusion, Submodule.map_top, Submodule.range_subtype] @@ -238,9 +238,9 @@ theorem LieAlgebra.isEngelian_of_isNoetherian [IsNoetherian R L] : LieAlgebra.Is LieSubmodule.nontrivial_iff_ne_bot R K] have : Nontrivial (L' ⧸ K.toLieSubmodule) := by replace hK₂ : K.toLieSubmodule ≠ ⊤ := by - rwa [Ne, ← LieSubmodule.coe_toSubmodule_inj, K.coe_toLieSubmodule, - LieSubmodule.top_coeSubmodule, ← LieSubalgebra.top_coe_submodule, - K.coe_to_submodule_inj] + rwa [Ne, ← LieSubmodule.toSubmodule_inj, K.coe_toLieSubmodule, + LieSubmodule.top_toSubmodule, ← LieSubalgebra.top_toSubmodule, + K.toSubmodule_inj] exact Submodule.Quotient.nontrivial_of_lt_top _ hK₂.lt_top have : LieModule.IsNilpotent R K (L' ⧸ K.toLieSubmodule) := by -- Porting note: was refine' hK₁ _ fun x => _ @@ -250,9 +250,9 @@ theorem LieAlgebra.isEngelian_of_isNoetherian [IsNoetherian R L] : LieAlgebra.Is apply Module.End.IsNilpotent.mapQ ?_ hx -- Porting note: mathlib3 solved this on its own with `submodule.mapq_linear._proof_5` intro X HX - simp only [LieSubalgebra.coe_toLieSubmodule, LieSubalgebra.mem_coe_submodule] at HX + simp only [LieSubalgebra.coe_toLieSubmodule, LieSubalgebra.mem_toSubmodule] at HX simp only [LieSubalgebra.coe_toLieSubmodule, Submodule.mem_comap, ad_apply, - LieSubalgebra.mem_coe_submodule] + LieSubalgebra.mem_toSubmodule] exact LieSubalgebra.lie_mem K x.prop HX exact nontrivial_max_triv_of_isNilpotent R K (L' ⧸ K.toLieSubmodule) haveI _i5 : IsNoetherian R L' := by @@ -260,7 +260,7 @@ theorem LieAlgebra.isEngelian_of_isNoetherian [IsNoetherian R L] : LieAlgebra.Is -- isNoetherian_of_surjective L _ (LinearMap.range_rangeRestrict (toEnd R L M)) -- abusing the relation between `LieHom.rangeRestrict` and `LinearMap.rangeRestrict` refine isNoetherian_of_surjective L (LieHom.rangeRestrict (toEnd R L M)) ?_ - simp only [LieHom.range_coeSubmodule, LieHom.coe_toLinearMap, + simp only [LieHom.range_toSubmodule, LieHom.coe_toLinearMap, LinearMap.range_eq_top] exact LieHom.surjective_rangeRestrict (toEnd R L M) obtain ⟨K, hK₁, hK₂⟩ := (LieSubalgebra.wellFoundedGT_of_noetherian R L').wf.has_min s hs diff --git a/Mathlib/Algebra/Lie/EngelSubalgebra.lean b/Mathlib/Algebra/Lie/EngelSubalgebra.lean index 6720922dae1f2..e011fcd6e98b4 100644 --- a/Mathlib/Algebra/Lie/EngelSubalgebra.lean +++ b/Mathlib/Algebra/Lie/EngelSubalgebra.lean @@ -122,7 +122,7 @@ lemma normalizer_eq_self_of_engel_le [IsArtinian R L] apply le_sup_of_le_left rw [Submodule.map_le_iff_le_comap] intro y hy - simp only [Submodule.mem_comap, mem_engel_iff, mem_coe_submodule] + simp only [Submodule.mem_comap, mem_engel_iff, mem_toSubmodule] use k+1 clear hk; revert hy generalize k+1 = k @@ -135,7 +135,7 @@ lemma normalizer_eq_self_of_engel_le [IsArtinian R L] apply le_sup_of_le_right rw [Submodule.map_le_iff_le_comap] rintro _ ⟨y, rfl⟩ - simp only [pow_succ', LinearMap.mul_apply, Submodule.mem_comap, mem_coe_submodule] + simp only [pow_succ', LinearMap.mul_apply, Submodule.mem_comap, mem_toSubmodule] apply aux₁ simp only [Submodule.coe_subtype, SetLike.coe_mem] diff --git a/Mathlib/Algebra/Lie/IdealOperations.lean b/Mathlib/Algebra/Lie/IdealOperations.lean index 4fd1e86e29fbf..aae8ed34539a5 100644 --- a/Mathlib/Algebra/Lie/IdealOperations.lean +++ b/Mathlib/Algebra/Lie/IdealOperations.lean @@ -61,7 +61,7 @@ theorem comap_map_eq (hf : f.ker = ⊥) : comap f (map f N) = N := by @[simp] theorem map_comap_incl : map N.incl (comap N.incl N') = N ⊓ N' := by - rw [← coe_toSubmodule_inj] + rw [← toSubmodule_inj] exact (N : Submodule R M).map_comap_subtype N' variable [LieAlgebra R L] [LieModule R L M₂] (I J : LieIdeal R L) @@ -201,7 +201,7 @@ theorem inf_lie : ⁅I ⊓ J, N⁆ ≤ ⁅I, N⁆ ⊓ ⁅J, N⁆ := by apply mono_lie_left <;> [exact inf_le_left; exact inf_le_right] theorem map_bracket_eq [LieModule R L M] : map f ⁅I, N⁆ = ⁅I, map f N⁆ := by - rw [← coe_toSubmodule_inj, coeSubmodule_map, lieIdeal_oper_eq_linear_span, + rw [← toSubmodule_inj, toSubmodule_map, lieIdeal_oper_eq_linear_span, lieIdeal_oper_eq_linear_span, Submodule.map_span] congr ext m @@ -244,7 +244,7 @@ theorem map_bracket_le {I₁ I₂ : LieIdeal R L} : map f ⁅I₁, I₂⁆ ≤ theorem map_bracket_eq {I₁ I₂ : LieIdeal R L} (h : Function.Surjective f) : map f ⁅I₁, I₂⁆ = ⁅map f I₁, map f I₂⁆ := by suffices ⁅map f I₁, map f I₂⁆ ≤ map f ⁅I₁, I₂⁆ by exact le_antisymm (map_bracket_le f) this - rw [← LieSubmodule.coeSubmodule_le_coeSubmodule, coe_map_of_surjective h, + rw [← LieSubmodule.toSubmodule_le_toSubmodule, coe_map_of_surjective h, LieSubmodule.lieIdeal_oper_eq_linear_span, LieSubmodule.lieIdeal_oper_eq_linear_span, LinearMap.map_span] apply Submodule.span_mono @@ -266,8 +266,8 @@ theorem map_comap_incl {I₁ I₂ : LieIdeal R L} : map I₁.incl (comap I₁.in theorem comap_bracket_eq {J₁ J₂ : LieIdeal R L'} (h : f.IsIdealMorphism) : comap f ⁅f.idealRange ⊓ J₁, f.idealRange ⊓ J₂⁆ = ⁅comap f J₁, comap f J₂⁆ ⊔ f.ker := by - rw [← LieSubmodule.coe_toSubmodule_inj, comap_coeSubmodule, - LieSubmodule.sup_coe_toSubmodule, f.ker_coeSubmodule, ← Submodule.comap_map_eq, + rw [← LieSubmodule.toSubmodule_inj, comap_toSubmodule, + LieSubmodule.sup_toSubmodule, f.ker_toSubmodule, ← Submodule.comap_map_eq, LieSubmodule.lieIdeal_oper_eq_linear_span, LieSubmodule.lieIdeal_oper_eq_linear_span, LinearMap.map_span] congr; simp only [LieHom.coe_toLinearMap, Set.mem_setOf_eq]; ext y diff --git a/Mathlib/Algebra/Lie/InvariantForm.lean b/Mathlib/Algebra/Lie/InvariantForm.lean index e95fdc8b0f1bf..0722c41605ea8 100644 --- a/Mathlib/Algebra/Lie/InvariantForm.lean +++ b/Mathlib/Algebra/Lie/InvariantForm.lean @@ -72,7 +72,7 @@ def orthogonal (hΦ_inv : Φ.lieInvariant L) (N : LieSubmodule R L M) : LieSubmo suffices (∀ n ∈ N, Φ n y = 0) → ∀ n ∈ N, Φ n ⁅x, y⁆ = 0 by simpa only [LinearMap.BilinForm.isOrtho_def, -- and some default simp lemmas AddSubsemigroup.mem_carrier, AddSubmonoid.mem_toSubsemigroup, Submodule.mem_toAddSubmonoid, - LinearMap.BilinForm.mem_orthogonal_iff, LieSubmodule.mem_coeSubmodule] + LinearMap.BilinForm.mem_orthogonal_iff, LieSubmodule.mem_toSubmodule] intro H a ha rw [← neg_eq_zero, ← hΦ_inv] exact H _ <| N.lie_mem ha @@ -125,31 +125,34 @@ variable (hL : ∀ I : LieIdeal K L, IsAtom I → ¬IsLieAbelian I) include hΦ_nondeg hΦ_refl hL open Module Submodule in -lemma orthogonal_isCompl_coe_submodule (I : LieIdeal K L) (hI : IsAtom I) : +lemma orthogonal_isCompl_toSubmodule (I : LieIdeal K L) (hI : IsAtom I) : IsCompl I.toSubmodule (orthogonal Φ hΦ_inv I).toSubmodule := by rw [orthogonal_toSubmodule, LinearMap.BilinForm.isCompl_orthogonal_iff_disjoint hΦ_refl, - ← orthogonal_toSubmodule _ hΦ_inv, ← LieSubmodule.disjoint_iff_coe_toSubmodule] + ← orthogonal_toSubmodule _ hΦ_inv, ← LieSubmodule.disjoint_iff_toSubmodule] exact orthogonal_disjoint Φ hΦ_nondeg hΦ_inv hL I hI +@[deprecated (since := "2024-12-30")] +alias orthogonal_isCompl_coe_submodule := orthogonal_isCompl_toSubmodule + open Module Submodule in lemma orthogonal_isCompl (I : LieIdeal K L) (hI : IsAtom I) : IsCompl I (orthogonal Φ hΦ_inv I) := by - rw [LieSubmodule.isCompl_iff_coe_toSubmodule] - exact orthogonal_isCompl_coe_submodule Φ hΦ_nondeg hΦ_inv hΦ_refl hL I hI + rw [LieSubmodule.isCompl_iff_toSubmodule] + exact orthogonal_isCompl_toSubmodule Φ 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] - exact orthogonal_isCompl_coe_submodule Φ hΦ_nondeg hΦ_inv hΦ_refl hL I hI + exact orthogonal_isCompl_toSubmodule Φ hΦ_nondeg hΦ_inv hΦ_refl hL I hI lemma restrict_orthogonal_nondegenerate (I : LieIdeal K L) (hI : IsAtom I) : (Φ.restrict (orthogonal Φ hΦ_inv I)).Nondegenerate := by rw [LinearMap.BilinForm.restrict_nondegenerate_iff_isCompl_orthogonal hΦ_refl] - simp only [LieIdeal.coe_to_lieSubalgebra_to_submodule, orthogonal_toSubmodule, + simp only [LieIdeal.coe_toLieSubalgebra_toSubmodule, orthogonal_toSubmodule, LinearMap.BilinForm.orthogonal_orthogonal hΦ_nondeg hΦ_refl] - exact (orthogonal_isCompl_coe_submodule Φ hΦ_nondeg hΦ_inv hΦ_refl hL I hI).symm + exact (orthogonal_isCompl_toSubmodule Φ hΦ_nondeg hΦ_inv hΦ_refl hL I hI).symm open Module Submodule in lemma atomistic : ∀ I : LieIdeal K L, sSup {J : LieIdeal K L | IsAtom J ∧ J ≤ I} = I := by diff --git a/Mathlib/Algebra/Lie/Nilpotent.lean b/Mathlib/Algebra/Lie/Nilpotent.lean index 96c142a2e09b1..87bf4540070a5 100644 --- a/Mathlib/Algebra/Lie/Nilpotent.lean +++ b/Mathlib/Algebra/Lie/Nilpotent.lean @@ -428,7 +428,7 @@ theorem coe_lcs_range_toEnd_eq (k : ℕ) : | zero => simp | succ k ih => simp only [lowerCentralSeries_succ, LieSubmodule.lieIdeal_oper_eq_linear_span', ← - (lowerCentralSeries R (toEnd R L M).range M k).mem_coeSubmodule, ih] + (lowerCentralSeries R (toEnd R L M).range M k).mem_toSubmodule, ih] congr ext m constructor @@ -442,7 +442,7 @@ theorem coe_lcs_range_toEnd_eq (k : ℕ) : theorem isNilpotent_range_toEnd_iff : IsNilpotent R (toEnd R L M).range M ↔ IsNilpotent R L M := by constructor <;> rintro ⟨k, hk⟩ <;> use k <;> - rw [← LieSubmodule.coe_toSubmodule_inj] at hk ⊢ <;> + rw [← LieSubmodule.toSubmodule_inj] at hk ⊢ <;> simpa using hk end LieModule @@ -550,13 +550,13 @@ theorem Function.Surjective.lieModule_lcs_map_eq (k : ℕ) : suffices g '' {m | ∃ (x : L) (n : _), n ∈ lowerCentralSeries R L M k ∧ ⁅x, n⁆ = m} = {m | ∃ (x : L₂) (n : _), n ∈ lowerCentralSeries R L M k ∧ ⁅x, g n⁆ = m} by - simp only [← LieSubmodule.mem_coeSubmodule] at this + simp only [← LieSubmodule.mem_toSubmodule] at this -- Porting note: was - -- simp [← LieSubmodule.mem_coeSubmodule, ← ih, LieSubmodule.lieIdeal_oper_eq_linear_span', + -- simp [← LieSubmodule.mem_toSubmodule, ← ih, LieSubmodule.lieIdeal_oper_eq_linear_span', -- Submodule.map_span, -Submodule.span_image, this, - -- -LieSubmodule.mem_coeSubmodule] + -- -LieSubmodule.mem_toSubmodule] simp_rw [lowerCentralSeries_succ, LieSubmodule.lieIdeal_oper_eq_linear_span', - Submodule.map_span, LieSubmodule.mem_top, true_and, ← LieSubmodule.mem_coeSubmodule, this, + Submodule.map_span, LieSubmodule.mem_top, true_and, ← LieSubmodule.mem_toSubmodule, this, ← ih, Submodule.mem_map, exists_exists_and_eq_and] ext m₂ constructor @@ -570,7 +570,7 @@ 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 - rw [← LieSubmodule.coe_toSubmodule_inj] at hk ⊢ + rw [← LieSubmodule.toSubmodule_inj] at hk ⊢ simp [← hf.lieModule_lcs_map_eq hg hfg k, hk] theorem Equiv.lieModule_isNilpotent_iff (f : L ≃ₗ⁅R⁆ L₂) (g : M ≃ₗ[R] M₂) @@ -580,7 +580,7 @@ theorem Equiv.lieModule_isNilpotent_iff (f : L ≃ₗ⁅R⁆ L₂) (g : M ≃ₗ exact f.surjective.lieModuleIsNilpotent hg hfg · have hg : Surjective (g.symm : M₂ →ₗ[R] M) := g.symm.surjective refine f.symm.surjective.lieModuleIsNilpotent hg fun x m => ?_ - rw [LinearEquiv.coe_coe, LieEquiv.coe_to_lieHom, ← g.symm_apply_apply ⁅f.symm x, g.symm m⁆, ← + rw [LinearEquiv.coe_coe, LieEquiv.coe_toLieHom, ← g.symm_apply_apply ⁅f.symm x, g.symm m⁆, ← hfg, f.apply_symm_apply, g.apply_symm_apply] @[simp] @@ -637,18 +637,18 @@ theorem coe_lowerCentralSeries_ideal_quot_eq {I : LieIdeal R L} (k : ℕ) : LieSubmodule.toSubmodule (lowerCentralSeries R (L ⧸ I) (L ⧸ I) k) := by induction k with | zero => - simp only [LieModule.lowerCentralSeries_zero, LieSubmodule.top_coeSubmodule, - LieIdeal.top_coe_lieSubalgebra, LieSubalgebra.top_coe_submodule] + simp only [LieModule.lowerCentralSeries_zero, LieSubmodule.top_toSubmodule, + LieIdeal.top_toLieSubalgebra, LieSubalgebra.top_toSubmodule] | succ k ih => simp only [LieModule.lowerCentralSeries_succ, LieSubmodule.lieIdeal_oper_eq_linear_span] congr ext x constructor · rintro ⟨⟨y, -⟩, ⟨z, hz⟩, rfl : ⁅y, z⁆ = x⟩ - rw [← LieSubmodule.mem_coeSubmodule, ih, LieSubmodule.mem_coeSubmodule] at hz + rw [← LieSubmodule.mem_toSubmodule, ih, LieSubmodule.mem_toSubmodule] at hz exact ⟨⟨LieSubmodule.Quotient.mk y, LieSubmodule.mem_top _⟩, ⟨z, hz⟩, rfl⟩ · rintro ⟨⟨⟨y⟩, -⟩, ⟨z, hz⟩, rfl : ⁅y, z⁆ = x⟩ - rw [← LieSubmodule.mem_coeSubmodule, ← ih, LieSubmodule.mem_coeSubmodule] at hz + rw [← LieSubmodule.mem_toSubmodule, ← ih, LieSubmodule.mem_toSubmodule] at hz exact ⟨⟨y, LieSubmodule.mem_top _⟩, ⟨z, hz⟩, rfl⟩ /-- Note that the below inequality can be strict. For example the ideal of strictly-upper-triangular @@ -671,7 +671,7 @@ theorem LieAlgebra.nilpotent_of_nilpotent_quotient {I : LieIdeal R L} (h₁ : I exact LieModule.nilpotentOfNilpotentQuotient R L L h₁ this obtain ⟨k, hk⟩ := h₂ use k - simp [← LieSubmodule.coe_toSubmodule_inj, coe_lowerCentralSeries_ideal_quot_eq, hk] + simp [← LieSubmodule.toSubmodule_inj, coe_lowerCentralSeries_ideal_quot_eq, hk] theorem LieAlgebra.non_trivial_center_of_isNilpotent [Nontrivial L] [IsNilpotent R L] : Nontrivial <| center R L := @@ -773,7 +773,7 @@ theorem coe_lcs_eq [LieModule R L M] : | zero => simp | succ k ih => simp_rw [lowerCentralSeries_succ, lcs_succ, LieSubmodule.lieIdeal_oper_eq_linear_span', ← - (I.lcs M k).mem_coeSubmodule, ih, LieSubmodule.mem_coeSubmodule, LieSubmodule.mem_top, + (I.lcs M k).mem_toSubmodule, ih, LieSubmodule.mem_toSubmodule, LieSubmodule.mem_top, true_and, (I : LieSubalgebra R L).coe_bracket_of_module] congr ext m diff --git a/Mathlib/Algebra/Lie/Normalizer.lean b/Mathlib/Algebra/Lie/Normalizer.lean index 233839a1ce04d..cf2530b63ccb0 100644 --- a/Mathlib/Algebra/Lie/Normalizer.lean +++ b/Mathlib/Algebra/Lie/Normalizer.lean @@ -147,7 +147,7 @@ theorem lie_mem_sup_of_mem_normalizer {x y z : L} (hx : x ∈ H.normalizer) (hy obtain ⟨t, rfl⟩ := Submodule.mem_span_singleton.mp hu₁ obtain ⟨s, rfl⟩ := Submodule.mem_span_singleton.mp hu₂ apply Submodule.mem_sup_right - simp only [LieSubalgebra.mem_coe_submodule, smul_lie, add_lie, zero_add, lie_add, smul_zero, + simp only [LieSubalgebra.mem_toSubmodule, smul_lie, add_lie, zero_add, lie_add, smul_zero, lie_smul, lie_self] refine H.add_mem (H.smul_mem s ?_) (H.add_mem (H.smul_mem t ?_) (H.lie_mem hv hw)) exacts [(H.mem_normalizer_iff' x).mp hx v hv, (H.mem_normalizer_iff x).mp hx w hw] @@ -172,7 +172,7 @@ theorem normalizer_eq_self_iff : refine ⟨fun h => ?_, fun h => le_antisymm ?_ H.le_normalizer⟩ · rintro ⟨x⟩ hx suffices x ∈ H by rwa [Submodule.Quotient.quot_mk_eq_mk, Submodule.Quotient.mk_eq_zero, - coe_toLieSubmodule, mem_coe_submodule] + coe_toLieSubmodule, mem_toSubmodule] rw [← h, H.mem_normalizer_iff'] intro y hy replace hx : ⁅_, LieSubmodule.Quotient.mk' _ x⁆ = 0 := hx ⟨y, hy⟩ diff --git a/Mathlib/Algebra/Lie/OfAssociative.lean b/Mathlib/Algebra/Lie/OfAssociative.lean index 10d1f4dcca458..446f21a24ea94 100644 --- a/Mathlib/Algebra/Lie/OfAssociative.lean +++ b/Mathlib/Algebra/Lie/OfAssociative.lean @@ -398,7 +398,7 @@ of the endomorphism `ad(x)` of `L` by `e` is the endomorphism `ad(e x)` of `L'`. @[simp] lemma conj_ad_apply (e : L ≃ₗ⁅R⁆ L') (x : L) : LinearEquiv.conj e (ad R L x) = ad R L' (e x) := by ext y' - rw [LinearEquiv.conj_apply_apply, ad_apply, ad_apply, coe_to_linearEquiv, map_lie, - ← coe_to_linearEquiv, LinearEquiv.apply_symm_apply] + rw [LinearEquiv.conj_apply_apply, ad_apply, ad_apply, coe_toLinearEquiv, map_lie, + ← coe_toLinearEquiv, LinearEquiv.apply_symm_apply] end LieAlgebra diff --git a/Mathlib/Algebra/Lie/Semisimple/Basic.lean b/Mathlib/Algebra/Lie/Semisimple/Basic.lean index 791f324394561..2bcaab54edd7e 100644 --- a/Mathlib/Algebra/Lie/Semisimple/Basic.lean +++ b/Mathlib/Algebra/Lie/Semisimple/Basic.lean @@ -140,7 +140,7 @@ lemma isSimple_of_isAtom (I : LieIdeal R L) (hI : IsAtom I) : IsSimple R I where Submodule.mem_toAddSubmonoid] apply add_mem -- Now `⁅a, y⁆ ∈ J` since `a ∈ I`, `y ∈ J`, and `J` is an ideal of `I`. - · simp only [Submodule.mem_map, LieSubmodule.mem_coeSubmodule, Subtype.exists] + · simp only [Submodule.mem_map, LieSubmodule.mem_toSubmodule, Subtype.exists] erw [Submodule.coe_subtype] simp only [exists_and_right, exists_eq_right, ha, lie_mem_left, exists_true_left] exact lie_mem_right R I J ⟨a, ha⟩ y hy @@ -159,8 +159,8 @@ lemma isSimple_of_isAtom (I : LieIdeal R L) (hI : IsAtom I) : IsSimple R I where intro x hx suffices x ∈ J → x = 0 from this hx have := @this x.1 - simp only [LieIdeal.incl_coe, LieIdeal.coe_to_lieSubalgebra_to_submodule, - LieSubmodule.mem_mk_iff', Submodule.mem_map, LieSubmodule.mem_coeSubmodule, Subtype.exists, + simp only [LieIdeal.incl_coe, LieIdeal.coe_toLieSubalgebra_toSubmodule, + LieSubmodule.mem_mk_iff', Submodule.mem_map, LieSubmodule.mem_toSubmodule, Subtype.exists, LieSubmodule.mem_bot, ZeroMemClass.coe_eq_zero, forall_exists_index, and_imp, J'] at this exact fun _ ↦ this (↑x) x.property hx rfl -- We need to show that `J = ⊥`. diff --git a/Mathlib/Algebra/Lie/Subalgebra.lean b/Mathlib/Algebra/Lie/Subalgebra.lean index 4c596fc5e43c6..661062dd102f7 100644 --- a/Mathlib/Algebra/Lie/Subalgebra.lean +++ b/Mathlib/Algebra/Lie/Subalgebra.lean @@ -147,9 +147,11 @@ theorem mem_mk_iff (S : Set L) (h₁ h₂ h₃ h₄) {x : L} : Iff.rfl @[simp] -theorem mem_coe_submodule {x : L} : x ∈ (L' : Submodule R L) ↔ x ∈ L' := +theorem mem_toSubmodule {x : L} : x ∈ (L' : Submodule R L) ↔ x ∈ L' := Iff.rfl +@[deprecated (since := "2024-12-30")] alias mem_coe_submodule := mem_toSubmodule + theorem mem_coe {x : L} : x ∈ (L' : Set L) ↔ x ∈ L' := Iff.rfl @@ -175,11 +177,13 @@ theorem mk_coe (S : Set L) (h₁ h₂ h₃ h₄) : ((⟨⟨⟨⟨S, h₁⟩, h₂⟩, h₃⟩, h₄⟩ : LieSubalgebra R L) : Set L) = S := rfl -theorem coe_to_submodule_mk (p : Submodule R L) (h) : +theorem toSubmodule_mk (p : Submodule R L) (h) : (({ p with lie_mem' := h } : LieSubalgebra R L) : Submodule R L) = p := by cases p rfl +@[deprecated (since := "2024-12-30")] alias coe_to_submodule_mk := toSubmodule_mk + theorem coe_injective : Function.Injective ((↑) : LieSubalgebra R L → Set L) := SetLike.coe_injective @@ -187,22 +191,28 @@ theorem coe_injective : Function.Injective ((↑) : LieSubalgebra R L → Set L) theorem coe_set_eq (L₁' L₂' : LieSubalgebra R L) : (L₁' : Set L) = L₂' ↔ L₁' = L₂' := SetLike.coe_set_eq -theorem to_submodule_injective : Function.Injective ((↑) : LieSubalgebra R L → Submodule R L) := +theorem toSubmodule_injective : Function.Injective ((↑) : LieSubalgebra R L → Submodule R L) := fun L₁' L₂' h ↦ by rw [SetLike.ext'_iff] at h rw [← coe_set_eq] exact h +@[deprecated (since := "2024-12-30")] alias to_submodule_injective := toSubmodule_injective + @[simp] -theorem coe_to_submodule_inj (L₁' L₂' : LieSubalgebra R L) : +theorem toSubmodule_inj (L₁' L₂' : LieSubalgebra R L) : (L₁' : Submodule R L) = (L₂' : Submodule R L) ↔ L₁' = L₂' := - to_submodule_injective.eq_iff + toSubmodule_injective.eq_iff -@[deprecated (since := "2024-12-29")] alias coe_to_submodule_eq_iff := coe_to_submodule_inj +@[deprecated (since := "2024-12-30")] alias coe_to_submodule_inj := toSubmodule_inj -theorem coe_to_submodule : ((L' : Submodule R L) : Set L) = L' := +@[deprecated (since := "2024-12-29")] alias toSubmodule_eq_iff := toSubmodule_inj + +theorem coe_toSubmodule : ((L' : Submodule R L) : Set L) = L' := rfl +@[deprecated (since := "2024-12-30")] alias coe_to_submodule := coe_toSubmodule + section LieModule variable {M : Type w} [AddCommGroup M] [LieRingModule L M] @@ -329,7 +339,7 @@ variable (K K' : LieSubalgebra R L) (K₂ : LieSubalgebra R L₂) @[simp] theorem incl_range : K.incl.range = K := by - rw [← coe_to_submodule_inj] + rw [← toSubmodule_inj] exact (K : Submodule R L).range_subtype /-- The image of a Lie subalgebra under a Lie algebra morphism is a Lie subalgebra of the @@ -375,9 +385,12 @@ theorem le_def : K ≤ K' ↔ (K : Set L) ⊆ K' := Iff.rfl @[simp] -theorem coe_submodule_le_coe_submodule : (K : Submodule R L) ≤ K' ↔ K ≤ K' := +theorem toSubmodule_le_toSubmodule : (K : Submodule R L) ≤ K' ↔ K ≤ K' := Iff.rfl +@[deprecated (since := "2024-12-30")] +alias coe_submodule_le_coe_submodule := toSubmodule_le_toSubmodule + instance : Bot (LieSubalgebra R L) := ⟨0⟩ @@ -386,9 +399,11 @@ theorem bot_coe : ((⊥ : LieSubalgebra R L) : Set L) = {0} := rfl @[simp] -theorem bot_coe_submodule : ((⊥ : LieSubalgebra R L) : Submodule R L) = ⊥ := +theorem bot_toSubmodule : ((⊥ : LieSubalgebra R L) : Submodule R L) = ⊥ := rfl +@[deprecated (since := "2024-12-30")] alias bot_coe_submodule := bot_toSubmodule + @[simp] theorem mem_bot (x : L) : x ∈ (⊥ : LieSubalgebra R L) ↔ x = 0 := mem_singleton_iff @@ -401,9 +416,11 @@ theorem top_coe : ((⊤ : LieSubalgebra R L) : Set L) = univ := rfl @[simp] -theorem top_coe_submodule : ((⊤ : LieSubalgebra R L) : Submodule R L) = ⊤ := +theorem top_toSubmodule : ((⊤ : LieSubalgebra R L) : Submodule R L) = ⊤ := rfl +@[deprecated (since := "2024-12-30")] alias top_coe_submodule := top_toSubmodule + @[simp] theorem mem_top (x : L) : x ∈ (⊤ : LieSubalgebra R L) := mem_univ x @@ -431,13 +448,15 @@ theorem inf_coe : (↑(K ⊓ K') : Set L) = (K : Set L) ∩ (K' : Set L) := rfl @[simp] -theorem sInf_coe_to_submodule (S : Set (LieSubalgebra R L)) : +theorem sInf_toSubmodule (S : Set (LieSubalgebra R L)) : (↑(sInf S) : Submodule R L) = sInf {(s : Submodule R L) | s ∈ S} := rfl +@[deprecated (since := "2024-12-30")] alias sInf_coe_to_submodule := sInf_toSubmodule + @[simp] theorem sInf_coe (S : Set (LieSubalgebra R L)) : (↑(sInf S) : Set L) = ⋂ s ∈ S, (s : Set L) := by - rw [← coe_to_submodule, sInf_coe_to_submodule, Submodule.sInf_coe] + rw [← coe_toSubmodule, sInf_toSubmodule, Submodule.sInf_coe] ext x simp @@ -490,13 +509,15 @@ theorem add_eq_sup : K + K' = K ⊔ K' := rfl @[simp] -theorem inf_coe_to_submodule : +theorem inf_toSubmodule : (↑(K ⊓ K') : Submodule R L) = (K : Submodule R L) ⊓ (K' : Submodule R L) := rfl +@[deprecated (since := "2024-12-30")] alias inf_coe_to_submodule := inf_toSubmodule + @[simp] theorem mem_inf (x : L) : x ∈ K ⊓ K' ↔ x ∈ K ∧ x ∈ K' := by - rw [← mem_coe_submodule, ← mem_coe_submodule, ← mem_coe_submodule, inf_coe_to_submodule, + rw [← mem_toSubmodule, ← mem_toSubmodule, ← mem_toSubmodule, inf_toSubmodule, Submodule.mem_inf] theorem eq_bot_iff : K = ⊥ ↔ ∀ x : L, x ∈ K → x = 0 := by @@ -623,9 +644,9 @@ theorem coe_lieSpan_submodule_eq_iff {p : Submodule R L} : (lieSpan R L (p : Set L) : Submodule R L) = p ↔ ∃ K : LieSubalgebra R L, ↑K = p := by rw [p.exists_lieSubalgebra_coe_eq_iff]; constructor <;> intro h · intro x m hm - rw [← h, mem_coe_submodule] + rw [← h, mem_toSubmodule] exact lie_mem _ (subset_lieSpan hm) - · rw [← coe_to_submodule_mk p @h, coe_to_submodule, coe_to_submodule_inj, lieSpan_eq] + · rw [← toSubmodule_mk p @h, coe_toSubmodule, toSubmodule_inj, lieSpan_eq] variable (R L) diff --git a/Mathlib/Algebra/Lie/Submodule.lean b/Mathlib/Algebra/Lie/Submodule.lean index 5c39a968185a7..24345ff122879 100644 --- a/Mathlib/Algebra/Lie/Submodule.lean +++ b/Mathlib/Algebra/Lie/Submodule.lean @@ -78,12 +78,11 @@ instance (priority := high) coeSort : CoeSort (LieSubmodule R L M) (Type w) wher instance (priority := mid) coeSubmodule : CoeOut (LieSubmodule R L M) (Submodule R M) := ⟨toSubmodule⟩ - @[norm_cast] theorem coe_toSubmodule : ((N : Submodule R M) : Set M) = N := rfl --- `simp` can prove this after `mem_coeSubmodule` is added to the simp set, +-- `simp` can prove this after `mem_toSubmodule` is added to the simp set, -- but `dsimp` can't. @[simp, nolint simpNF] theorem mem_carrier {x : M} : x ∈ N.carrier ↔ x ∈ (N : Set M) := @@ -99,9 +98,11 @@ theorem mem_mk_iff' (p : Submodule R M) (h) {x : M} : Iff.rfl @[simp] -theorem mem_coeSubmodule {x : M} : x ∈ (N : Submodule R M) ↔ x ∈ N := +theorem mem_toSubmodule {x : M} : x ∈ (N : Submodule R M) ↔ x ∈ N := Iff.rfl +@[deprecated (since := "2024-12-30")] alias mem_coeSubmodule := mem_toSubmodule + theorem mem_coe {x : M} : x ∈ (N : Set M) ↔ x ∈ N := Iff.rfl @@ -118,22 +119,28 @@ theorem coe_toSet_mk (S : Set M) (h₁ h₂ h₃ h₄) : ((⟨⟨⟨⟨S, h₁⟩, h₂⟩, h₃⟩, h₄⟩ : LieSubmodule R L M) : Set M) = S := rfl -theorem coe_toSubmodule_mk (p : Submodule R M) (h) : +theorem toSubmodule_mk (p : Submodule R M) (h) : (({ p with lie_mem := h } : LieSubmodule R L M) : Submodule R M) = p := by cases p; rfl -theorem coeSubmodule_injective : +@[deprecated (since := "2024-12-30")] alias coe_toSubmodule_mk := toSubmodule_mk + +theorem toSubmodule_injective : Function.Injective (toSubmodule : LieSubmodule R L M → Submodule R M) := fun x y h ↦ by cases x; cases y; congr +@[deprecated (since := "2024-12-30")] alias coeSubmodule_injective := toSubmodule_injective + @[ext] theorem ext (h : ∀ m, m ∈ N ↔ m ∈ N') : N = N' := SetLike.ext h @[simp] -theorem coe_toSubmodule_inj : (N : Submodule R M) = (N' : Submodule R M) ↔ N = N' := - coeSubmodule_injective.eq_iff +theorem toSubmodule_inj : (N : Submodule R M) = (N' : Submodule R M) ↔ N = N' := + toSubmodule_injective.eq_iff + +@[deprecated (since := "2024-12-30")] alias coe_toSubmodule_inj := toSubmodule_inj -@[deprecated (since := "2024-12-29")] alias coe_toSubmodule_eq_iff := coe_toSubmodule_inj +@[deprecated (since := "2024-12-29")] alias toSubmodule_eq_iff := toSubmodule_inj /-- Copy of a `LieSubmodule` with a new `carrier` equal to the old one. Useful to fix definitional equalities. -/ @@ -200,7 +207,7 @@ instance instLieModule : LieModule R L N where 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_inj _ _).mp (Subsingleton.elim _ _)⟩ + ⟨⟨0⟩, fun _ ↦ (toSubmodule_inj _ _).mp (Subsingleton.elim _ _)⟩ end LieSubmodule @@ -225,14 +232,20 @@ instance : Coe (LieIdeal R L) (LieSubalgebra R L) := ⟨lieIdealSubalgebra R L⟩ @[simp] -theorem LieIdeal.coe_toSubalgebra (I : LieIdeal R L) : ((I : LieSubalgebra R L) : Set L) = I := +theorem LieIdeal.coe_toLieSubalgebra (I : LieIdeal R L) : ((I : LieSubalgebra R L) : Set L) = I := rfl +@[deprecated (since := "2024-12-30")] +alias LieIdeal.coe_toSubalgebra := LieIdeal.coe_toLieSubalgebra + @[simp] -theorem LieIdeal.coe_to_lieSubalgebra_to_submodule (I : LieIdeal R L) : +theorem LieIdeal.coe_toLieSubalgebra_toSubmodule (I : LieIdeal R L) : ((I : LieSubalgebra R L) : Submodule R L) = LieSubmodule.toSubmodule I := rfl +@[deprecated (since := "2024-12-30")] +alias LieIdeal.coe_to_lieSubalgebra_to_submodule := LieIdeal.coe_toLieSubalgebra_toSubmodule + /-- An ideal of `L` is a Lie subalgebra of `L`, so it is a Lie ring. -/ instance LieIdeal.lieRing (I : LieIdeal R L) : LieRing I := LieSubalgebra.lieRing R L ↑I @@ -288,7 +301,7 @@ theorem mem_toLieSubmodule (x : L) : x ∈ K.toLieSubmodule ↔ x ∈ K := theorem exists_lieIdeal_coe_eq_iff : (∃ I : LieIdeal R L, ↑I = K) ↔ ∀ x y : L, y ∈ K → ⁅x, y⁆ ∈ K := by - simp only [← coe_to_submodule_inj, LieIdeal.coe_to_lieSubalgebra_to_submodule, + simp only [← toSubmodule_inj, LieIdeal.coe_toLieSubalgebra_toSubmodule, Submodule.exists_lieSubmodule_coe_eq_iff L] exact Iff.rfl @@ -318,9 +331,12 @@ theorem coe_injective : Function.Injective ((↑) : LieSubmodule R L M → Set M SetLike.coe_injective @[simp, norm_cast] -theorem coeSubmodule_le_coeSubmodule : (N : Submodule R M) ≤ N' ↔ N ≤ N' := +theorem toSubmodule_le_toSubmodule : (N : Submodule R M) ≤ N' ↔ N ≤ N' := Iff.rfl +@[deprecated (since := "2024-12-30")] +alias coeSubmodule_le_coeSubmodule := toSubmodule_le_toSubmodule + instance : Bot (LieSubmodule R L M) := ⟨0⟩ @@ -332,16 +348,20 @@ theorem bot_coe : ((⊥ : LieSubmodule R L M) : Set M) = {0} := rfl @[simp] -theorem bot_coeSubmodule : ((⊥ : LieSubmodule R L M) : Submodule R M) = ⊥ := +theorem bot_toSubmodule : ((⊥ : LieSubmodule R L M) : Submodule R M) = ⊥ := rfl +@[deprecated (since := "2024-12-30")] alias bot_coeSubmodule := bot_toSubmodule + @[simp] -theorem coeSubmodule_eq_bot_iff : (N : Submodule R M) = ⊥ ↔ N = ⊥ := by - rw [← coe_toSubmodule_inj, bot_coeSubmodule] +theorem toSubmodule_eq_bot : (N : Submodule R M) = ⊥ ↔ N = ⊥ := by + rw [← toSubmodule_inj, bot_toSubmodule] + +@[deprecated (since := "2024-12-30")] alias coeSubmodule_eq_bot_iff := toSubmodule_eq_bot @[simp] theorem mk_eq_bot_iff {N : Submodule R M} {h} : (⟨N, h⟩ : LieSubmodule R L M) = ⊥ ↔ N = ⊥ := by - rw [← coe_toSubmodule_inj, bot_coeSubmodule] + rw [← toSubmodule_inj, bot_toSubmodule] @[simp] theorem mem_bot (x : M) : x ∈ (⊥ : LieSubmodule R L M) ↔ x = 0 := @@ -355,16 +375,20 @@ theorem top_coe : ((⊤ : LieSubmodule R L M) : Set M) = univ := rfl @[simp] -theorem top_coeSubmodule : ((⊤ : LieSubmodule R L M) : Submodule R M) = ⊤ := +theorem top_toSubmodule : ((⊤ : LieSubmodule R L M) : Submodule R M) = ⊤ := rfl +@[deprecated (since := "2024-12-30")] alias top_coeSubmodule := top_toSubmodule + @[simp] -theorem coeSubmodule_eq_top_iff : (N : Submodule R M) = ⊤ ↔ N = ⊤ := by - rw [← coe_toSubmodule_inj, top_coeSubmodule] +theorem toSubmodule_eq_top : (N : Submodule R M) = ⊤ ↔ N = ⊤ := by + rw [← toSubmodule_inj, top_toSubmodule] + +@[deprecated (since := "2024-12-30")] alias coeSubmodule_eq_top_iff := toSubmodule_eq_top @[simp] theorem mk_eq_top_iff {N : Submodule R M} {h} : (⟨N, h⟩ : LieSubmodule R L M) = ⊤ ↔ N = ⊤ := by - rw [← coe_toSubmodule_inj, top_coeSubmodule] + rw [← toSubmodule_inj, top_toSubmodule] @[simp] theorem mem_top (x : M) : x ∈ (⊤ : LieSubmodule R L M) := @@ -388,30 +412,38 @@ theorem inf_coe : (↑(N ⊓ N') : Set M) = ↑N ∩ ↑N' := rfl @[norm_cast, simp] -theorem inf_coe_toSubmodule : +theorem inf_toSubmodule : (↑(N ⊓ N') : Submodule R M) = (N : Submodule R M) ⊓ (N' : Submodule R M) := rfl +@[deprecated (since := "2024-12-30")] alias inf_coe_toSubmodule := inf_toSubmodule + @[simp] -theorem sInf_coe_toSubmodule (S : Set (LieSubmodule R L M)) : +theorem sInf_toSubmodule (S : Set (LieSubmodule R L M)) : (↑(sInf S) : Submodule R M) = sInf {(s : Submodule R M) | s ∈ S} := rfl -theorem sInf_coe_toSubmodule' (S : Set (LieSubmodule R L M)) : +@[deprecated (since := "2024-12-30")] alias sInf_coe_toSubmodule := sInf_toSubmodule + +theorem sInf_toSubmodule_eq_iInf (S : Set (LieSubmodule R L M)) : (↑(sInf S) : Submodule R M) = ⨅ N ∈ S, (N : Submodule R M) := by - rw [sInf_coe_toSubmodule, ← Set.image, sInf_image] + rw [sInf_toSubmodule, ← Set.image, sInf_image] + +@[deprecated (since := "2024-12-30")] alias sInf_coe_toSubmodule' := sInf_toSubmodule_eq_iInf @[simp] -theorem iInf_coe_toSubmodule {ι} (p : ι → LieSubmodule R L M) : +theorem iInf_toSubmodule {ι} (p : ι → LieSubmodule R L M) : (↑(⨅ i, p i) : Submodule R M) = ⨅ i, (p i : Submodule R M) := by - rw [iInf, sInf_coe_toSubmodule]; ext; simp + rw [iInf, sInf_toSubmodule]; ext; simp + +@[deprecated (since := "2024-12-30")] alias iInf_coe_toSubmodule := iInf_toSubmodule @[simp] theorem sInf_coe (S : Set (LieSubmodule R L M)) : (↑(sInf S) : Set M) = ⋂ s ∈ S, (s : Set M) := by - rw [← LieSubmodule.coe_toSubmodule, sInf_coe_toSubmodule, Submodule.sInf_coe] + rw [← LieSubmodule.coe_toSubmodule, sInf_toSubmodule, Submodule.sInf_coe] ext m simp only [mem_iInter, mem_setOf_eq, forall_apply_eq_imp_iff₂, exists_imp, - and_imp, SetLike.mem_coe, mem_coeSubmodule] + and_imp, SetLike.mem_coe, mem_toSubmodule] @[simp] theorem iInf_coe {ι} (p : ι → LieSubmodule R L M) : (↑(⨅ i, p i) : Set M) = ⋂ i, ↑(p i) := by @@ -455,28 +487,36 @@ instance : SupSet (LieSubmodule R L M) where exact le_sSup ⟨p, hp, rfl⟩ } @[norm_cast, simp] -theorem sup_coe_toSubmodule : +theorem sup_toSubmodule : (↑(N ⊔ N') : Submodule R M) = (N : Submodule R M) ⊔ (N' : Submodule R M) := by rfl +@[deprecated (since := "2024-12-30")] alias sup_coe_toSubmodule := sup_toSubmodule + @[simp] -theorem sSup_coe_toSubmodule (S : Set (LieSubmodule R L M)) : +theorem sSup_toSubmodule (S : Set (LieSubmodule R L M)) : (↑(sSup S) : Submodule R M) = sSup {(s : Submodule R M) | s ∈ S} := rfl -theorem sSup_coe_toSubmodule' (S : Set (LieSubmodule R L M)) : +@[deprecated (since := "2024-12-30")] alias sSup_coe_toSubmodule := sSup_toSubmodule + +theorem sSup_toSubmodule_eq_iSup (S : Set (LieSubmodule R L M)) : (↑(sSup S) : Submodule R M) = ⨆ N ∈ S, (N : Submodule R M) := by - rw [sSup_coe_toSubmodule, ← Set.image, sSup_image] + rw [sSup_toSubmodule, ← Set.image, sSup_image] + +@[deprecated (since := "2024-12-30")] alias sSup_coe_toSubmodule' := sSup_toSubmodule_eq_iSup @[simp] -theorem iSup_coe_toSubmodule {ι} (p : ι → LieSubmodule R L M) : +theorem iSup_toSubmodule {ι} (p : ι → LieSubmodule R L M) : (↑(⨆ i, p i) : Submodule R M) = ⨆ i, (p i : Submodule R M) := by - rw [iSup, sSup_coe_toSubmodule]; ext; simp [Submodule.mem_sSup, Submodule.mem_iSup] + rw [iSup, sSup_toSubmodule]; ext; simp [Submodule.mem_sSup, Submodule.mem_iSup] + +@[deprecated (since := "2024-12-30")] alias iSup_coe_toSubmodule := iSup_toSubmodule /-- The set of Lie submodules of a Lie module form a complete lattice. -/ instance : CompleteLattice (LieSubmodule R L M) := - { coeSubmodule_injective.completeLattice toSubmodule sup_coe_toSubmodule inf_coe_toSubmodule - sSup_coe_toSubmodule' sInf_coe_toSubmodule' rfl rfl with + { toSubmodule_injective.completeLattice toSubmodule sup_toSubmodule inf_toSubmodule + sSup_toSubmodule_eq_iSup sInf_toSubmodule_eq_iInf rfl rfl with toPartialOrder := SetLike.instPartialOrder } theorem mem_iSup_of_mem {ι} {b : M} {N : ι → LieSubmodule R L M} (i : ι) (h : b ∈ N i) : @@ -486,7 +526,7 @@ theorem mem_iSup_of_mem {ι} {b : M} {N : ι → LieSubmodule R L M} (i : ι) (h lemma iSup_induction {ι} (N : ι → LieSubmodule R L M) {C : M → Prop} {x : M} (hx : x ∈ ⨆ i, N i) (hN : ∀ i, ∀ y ∈ N i, C y) (h0 : C 0) (hadd : ∀ y z, C y → C z → C (y + z)) : C x := by - rw [← LieSubmodule.mem_coeSubmodule, LieSubmodule.iSup_coe_toSubmodule] at hx + rw [← LieSubmodule.mem_toSubmodule, LieSubmodule.iSup_toSubmodule] at hx exact Submodule.iSup_induction (C := C) (fun i ↦ (N i : Submodule R M)) hx hN h0 hadd @[elab_as_elim] @@ -502,30 +542,47 @@ theorem iSup_induction' {ι} (N : ι → LieSubmodule R L M) {C : (x : M) → (x · rintro ⟨_, Cx⟩ ⟨_, Cy⟩ exact ⟨_, hadd _ _ _ _ Cx Cy⟩ -theorem disjoint_iff_coe_toSubmodule : +-- TODO(Yaël): turn around +theorem disjoint_iff_toSubmodule : Disjoint N N' ↔ Disjoint (N : Submodule R M) (N' : Submodule R M) := by - rw [disjoint_iff, disjoint_iff, ← coe_toSubmodule_inj, inf_coe_toSubmodule, bot_coeSubmodule, + rw [disjoint_iff, disjoint_iff, ← toSubmodule_inj, inf_toSubmodule, bot_toSubmodule, ← disjoint_iff] -theorem codisjoint_iff_coe_toSubmodule : +@[deprecated (since := "2024-12-30")] alias disjoint_iff_coe_toSubmodule := disjoint_iff_toSubmodule + +theorem codisjoint_iff_toSubmodule : Codisjoint N N' ↔ Codisjoint (N : Submodule R M) (N' : Submodule R M) := by - rw [codisjoint_iff, codisjoint_iff, ← coe_toSubmodule_inj, sup_coe_toSubmodule, - top_coeSubmodule, ← codisjoint_iff] + rw [codisjoint_iff, codisjoint_iff, ← toSubmodule_inj, sup_toSubmodule, + top_toSubmodule, ← codisjoint_iff] -theorem isCompl_iff_coe_toSubmodule : +@[deprecated (since := "2024-12-30")] +alias codisjoint_iff_coe_toSubmodule := codisjoint_iff_toSubmodule + +theorem isCompl_iff_toSubmodule : IsCompl N N' ↔ IsCompl (N : Submodule R M) (N' : Submodule R M) := by - simp only [isCompl_iff, disjoint_iff_coe_toSubmodule, codisjoint_iff_coe_toSubmodule] + simp only [isCompl_iff, disjoint_iff_toSubmodule, codisjoint_iff_toSubmodule] + +@[deprecated (since := "2024-12-30")] alias isCompl_iff_coe_toSubmodule := isCompl_iff_toSubmodule -theorem iSupIndep_iff_coe_toSubmodule {ι : Type*} {N : ι → LieSubmodule R L M} : +theorem iSupIndep_iff_toSubmodule {ι : Type*} {N : ι → LieSubmodule R L M} : iSupIndep N ↔ iSupIndep fun i ↦ (N i : Submodule R M) := by - simp [iSupIndep_def, disjoint_iff_coe_toSubmodule] + simp [iSupIndep_def, disjoint_iff_toSubmodule] + +@[deprecated (since := "2024-12-30")] +alias iSupIndep_iff_coe_toSubmodule := iSupIndep_iff_toSubmodule @[deprecated (since := "2024-11-24")] -alias independent_iff_coe_toSubmodule := iSupIndep_iff_coe_toSubmodule +alias independent_iff_toSubmodule := iSupIndep_iff_toSubmodule + +@[deprecated (since := "2024-12-30")] +alias independent_iff_coe_toSubmodule := independent_iff_toSubmodule -theorem iSup_eq_top_iff_coe_toSubmodule {ι : Sort*} {N : ι → LieSubmodule R L M} : +theorem iSup_eq_top_iff_toSubmodule {ι : Sort*} {N : ι → LieSubmodule R L M} : ⨆ i, N i = ⊤ ↔ ⨆ i, (N i : Submodule R M) = ⊤ := by - rw [← iSup_coe_toSubmodule, ← top_coeSubmodule (L := L), coe_toSubmodule_inj] + rw [← iSup_toSubmodule, ← top_toSubmodule (L := L), toSubmodule_inj] + +@[deprecated (since := "2024-12-30")] +alias iSup_eq_top_iff_coe_toSubmodule := iSup_eq_top_iff_toSubmodule instance : Add (LieSubmodule R L M) where add := max @@ -544,11 +601,11 @@ theorem add_eq_sup : N + N' = N ⊔ N' := @[simp] theorem mem_inf (x : M) : x ∈ N ⊓ N' ↔ x ∈ N ∧ x ∈ N' := by - rw [← mem_coeSubmodule, ← mem_coeSubmodule, ← mem_coeSubmodule, inf_coe_toSubmodule, + rw [← mem_toSubmodule, ← mem_toSubmodule, ← mem_toSubmodule, inf_toSubmodule, Submodule.mem_inf] theorem mem_sup (x : M) : x ∈ N ⊔ N' ↔ ∃ y ∈ N, ∃ z ∈ N', y + z = x := by - rw [← mem_coeSubmodule, sup_coe_toSubmodule, Submodule.mem_sup]; exact Iff.rfl + rw [← mem_toSubmodule, sup_toSubmodule, Submodule.mem_sup]; exact Iff.rfl nonrec theorem eq_bot_iff : N = ⊥ ↔ ∀ m : M, m ∈ N → m = 0 := by rw [eq_bot_iff]; exact Iff.rfl @@ -560,7 +617,7 @@ instance subsingleton_of_bot : Subsingleton (LieSubmodule R L (⊥ : LieSubmodul instance : IsModularLattice (LieSubmodule R L M) where sup_inf_le_assoc_of_le _ _ := by - simp only [← coeSubmodule_le_coeSubmodule, sup_coe_toSubmodule, inf_coe_toSubmodule] + simp only [← toSubmodule_le_toSubmodule, sup_toSubmodule, inf_toSubmodule] exact IsModularLattice.sup_inf_le_assoc_of_le _ variable (R L M) @@ -568,7 +625,7 @@ variable (R L M) /-- The natural functor that forgets the action of `L` as an order embedding. -/ @[simps] def toSubmodule_orderEmbedding : LieSubmodule R L M ↪o Submodule R M := { toFun := (↑) - inj' := coeSubmodule_injective + inj' := toSubmodule_injective map_rel_iff' := Iff.rfl } instance wellFoundedGT_of_noetherian [IsNoetherian R M] : WellFoundedGT (LieSubmodule R L M) := @@ -583,8 +640,8 @@ instance [IsArtinian R M] : IsAtomic (LieSubmodule R L M) := @[simp] theorem subsingleton_iff : Subsingleton (LieSubmodule R L M) ↔ Subsingleton M := have h : Subsingleton (LieSubmodule R L M) ↔ Subsingleton (Submodule R M) := by - rw [← subsingleton_iff_bot_eq_top, ← subsingleton_iff_bot_eq_top, ← coe_toSubmodule_inj, - top_coeSubmodule, bot_coeSubmodule] + rw [← subsingleton_iff_bot_eq_top, ← subsingleton_iff_bot_eq_top, ← toSubmodule_inj, + top_toSubmodule, bot_toSubmodule] h.trans <| Submodule.subsingleton_iff R @[simp] @@ -686,8 +743,8 @@ theorem lieSpan_eq : lieSpan R L (N : Set M) = N := theorem coe_lieSpan_submodule_eq_iff {p : Submodule R M} : (lieSpan R L (p : Set M) : Submodule R M) = p ↔ ∃ N : LieSubmodule R L M, ↑N = p := by rw [p.exists_lieSubmodule_coe_eq_iff L]; constructor <;> intro h - · intro x m hm; rw [← h, mem_coeSubmodule]; exact lie_mem _ (subset_lieSpan hm) - · rw [← coe_toSubmodule_mk p @h, coe_toSubmodule, coe_toSubmodule_inj, lieSpan_eq] + · intro x m hm; rw [← h, mem_toSubmodule]; exact lie_mem _ (subset_lieSpan hm) + · rw [← toSubmodule_mk p @h, coe_toSubmodule, toSubmodule_inj, lieSpan_eq] variable (R L M) @@ -728,13 +785,13 @@ lemma isCompactElement_lieSpan_singleton (m : M) : exact ⟨N, hN, by simpa⟩ replace hne : Nonempty s := Set.nonempty_coe_sort.mpr hne have := Submodule.coe_iSup_of_directed _ hdir.directed_val - simp_rw [← iSup_coe_toSubmodule, Set.iUnion_coe_set, coe_toSubmodule] at this + simp_rw [← iSup_toSubmodule, Set.iUnion_coe_set, coe_toSubmodule] at this rw [← this, SetLike.coe_set_eq, sSup_eq_iSup, iSup_subtype] @[simp] lemma sSup_image_lieSpan_singleton : sSup ((fun x ↦ lieSpan R L {x}) '' N) = N := by refine le_antisymm (sSup_le <| by simp) ?_ - simp_rw [← coeSubmodule_le_coeSubmodule, sSup_coe_toSubmodule, Set.mem_image, SetLike.mem_coe] + simp_rw [← toSubmodule_le_toSubmodule, sSup_toSubmodule, Set.mem_image, SetLike.mem_coe] refine fun m hm ↦ Submodule.mem_sSup.mpr fun N' hN' ↦ ?_ replace hN' : ∀ m ∈ N, lieSpan R L {m} ≤ N' := by simpa using hN' exact hN' _ hm (subset_lieSpan rfl) @@ -772,9 +829,11 @@ def map : LieSubmodule R L M' := @[simp] theorem coe_map : (N.map f : Set M') = f '' N := rfl @[simp] -theorem coeSubmodule_map : (N.map f : Submodule R M') = (N : Submodule R M).map (f : M →ₗ[R] M') := +theorem toSubmodule_map : (N.map f : Submodule R M') = (N : Submodule R M).map (f : M →ₗ[R] M') := rfl +@[deprecated (since := "2024-12-30")] alias coeSubmodule_map := toSubmodule_map + /-- A morphism of Lie modules `f : M → M'` pulls back Lie submodules of `M'` to Lie submodules of `M`. -/ def comap : LieSubmodule R L M := @@ -784,10 +843,12 @@ def comap : LieSubmodule R L M := apply N'.lie_mem h } @[simp] -theorem coeSubmodule_comap : +theorem toSubmodule_comap : (N'.comap f : Submodule R M) = (N' : Submodule R M').comap (f : M →ₗ[R] M') := rfl +@[deprecated (since := "2024-12-30")] alias coeSubmodule_comap := toSubmodule_comap + variable {f N N₂ N'} theorem map_le_iff_le_comap : map f N ≤ N' ↔ N ≤ comap f N' := @@ -832,12 +893,12 @@ theorem mem_comap {m : M} : m ∈ comap f N' ↔ f m ∈ N' := Iff.rfl theorem comap_incl_eq_top : N₂.comap N.incl = ⊤ ↔ N ≤ N₂ := by - rw [← LieSubmodule.coe_toSubmodule_inj, LieSubmodule.coeSubmodule_comap, LieSubmodule.incl_coe, - LieSubmodule.top_coeSubmodule, Submodule.comap_subtype_eq_top, coeSubmodule_le_coeSubmodule] + rw [← LieSubmodule.toSubmodule_inj, LieSubmodule.toSubmodule_comap, LieSubmodule.incl_coe, + LieSubmodule.top_toSubmodule, Submodule.comap_subtype_eq_top, toSubmodule_le_toSubmodule] theorem comap_incl_eq_bot : N₂.comap N.incl = ⊥ ↔ N ⊓ N₂ = ⊥ := by - simp only [← coe_toSubmodule_inj, coeSubmodule_comap, incl_coe, bot_coeSubmodule, - inf_coe_toSubmodule] + simp only [← toSubmodule_inj, toSubmodule_comap, incl_coe, bot_toSubmodule, + inf_toSubmodule] rw [← Submodule.disjoint_iff_comap_eq_bot, disjoint_iff] @[gcongr, mono] @@ -900,9 +961,11 @@ 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') @[simp] -theorem top_coe_lieSubalgebra : ((⊤ : LieIdeal R L) : LieSubalgebra R L) = ⊤ := +theorem top_toLieSubalgebra : ((⊤ : LieIdeal R L) : LieSubalgebra R L) = ⊤ := rfl +@[deprecated (since := "2024-12-30")] alias top_coe_lieSubalgebra := top_toLieSubalgebra + /-- A morphism of Lie algebras `f : L → L'` pushes forward Lie ideals of `L` to Lie ideals of `L'`. Note that unlike `LieSubmodule.map`, we must take the `lieSpan` of the image. Mathematically @@ -921,20 +984,24 @@ def comap : LieIdeal R L := suffices ⁅f x, f y⁆ ∈ J by simp only [AddSubsemigroup.mem_carrier, AddSubmonoid.mem_toSubsemigroup, Submodule.mem_toAddSubmonoid, Submodule.mem_comap, LieHom.coe_toLinearMap, LieHom.map_lie, - LieSubalgebra.mem_coe_submodule] + LieSubalgebra.mem_toSubmodule] exact this apply J.lie_mem h } @[simp] -theorem map_coeSubmodule (h : ↑(map f I) = f '' I) : +theorem map_toSubmodule (h : ↑(map f I) = f '' I) : LieSubmodule.toSubmodule (map f I) = (LieSubmodule.toSubmodule I).map (f : L →ₗ[R] L') := by rw [SetLike.ext'_iff, LieSubmodule.coe_toSubmodule, h, Submodule.map_coe]; rfl +@[deprecated (since := "2024-12-30")] alias map_coeSubmodule := map_toSubmodule + @[simp] -theorem comap_coeSubmodule : +theorem comap_toSubmodule : (LieSubmodule.toSubmodule (comap f J)) = (LieSubmodule.toSubmodule J).comap (f : L →ₗ[R] L') := rfl +@[deprecated (since := "2024-12-30")] alias comap_coeSubmodule := comap_toSubmodule + theorem map_le : map f I ≤ J ↔ f '' I ⊆ J := LieSubmodule.lieSpan_le @@ -1032,9 +1099,9 @@ theorem IsIdealMorphism.eq (hf : f.IsIdealMorphism) : f.idealRange = f.range := theorem isIdealMorphism_iff : f.IsIdealMorphism ↔ ∀ (x : L') (y : L), ∃ z : L, ⁅x, f y⁆ = f z := by simp only [isIdealMorphism_def, idealRange_eq_lieSpan_range, ← - LieSubalgebra.coe_to_submodule_inj, ← f.range.coe_to_submodule, - LieIdeal.coe_to_lieSubalgebra_to_submodule, LieSubmodule.coe_lieSpan_submodule_eq_iff, - LieSubalgebra.mem_coe_submodule, mem_range, exists_imp, + LieSubalgebra.toSubmodule_inj, ← f.range.coe_toSubmodule, + LieIdeal.coe_toLieSubalgebra_toSubmodule, LieSubmodule.coe_lieSpan_submodule_eq_iff, + LieSubalgebra.mem_toSubmodule, mem_range, exists_imp, Submodule.exists_lieSubmodule_coe_eq_iff] constructor · intro h x y; obtain ⟨z, hz⟩ := h x (f y) y rfl; use z; exact hz.symm @@ -1051,14 +1118,16 @@ theorem ker_le_comap : f.ker ≤ J.comap f := LieIdeal.comap_mono bot_le @[simp] -theorem ker_coeSubmodule : LieSubmodule.toSubmodule (ker f) = LinearMap.ker (f : L →ₗ[R] L') := +theorem ker_toSubmodule : LieSubmodule.toSubmodule (ker f) = LinearMap.ker (f : L →ₗ[R] L') := rfl +@[deprecated (since := "2024-12-30")] alias ker_coeSubmodule := ker_toSubmodule + variable {f} in @[simp] theorem mem_ker {x : L} : x ∈ ker f ↔ f x = 0 := show x ∈ LieSubmodule.toSubmodule (f.ker) ↔ _ by - simp only [ker_coeSubmodule, LinearMap.mem_ker, coe_toLinearMap] + simp only [ker_toSubmodule, LinearMap.mem_ker, coe_toLinearMap] theorem mem_idealRange (x : L) : f x ∈ idealRange f := by rw [idealRange_eq_map] @@ -1068,7 +1137,7 @@ theorem mem_idealRange (x : L) : f x ∈ idealRange f := by theorem mem_idealRange_iff (h : IsIdealMorphism f) {y : L'} : y ∈ idealRange f ↔ ∃ x : L, f x = y := by rw [f.isIdealMorphism_def] at h - rw [← LieSubmodule.mem_coe, ← LieIdeal.coe_toSubalgebra, h, f.range_coe, Set.mem_range] + rw [← LieSubmodule.mem_coe, ← LieIdeal.coe_toLieSubalgebra, h, f.range_coe, Set.mem_range] theorem le_ker_iff : I ≤ f.ker ↔ ∀ x, x ∈ I → f x = 0 := by constructor <;> intro h x hx @@ -1076,29 +1145,31 @@ theorem le_ker_iff : I ≤ f.ker ↔ ∀ x, x ∈ I → f x = 0 := by · rw [mem_ker]; apply h x hx theorem ker_eq_bot : f.ker = ⊥ ↔ Function.Injective f := by - rw [← LieSubmodule.coe_toSubmodule_inj, ker_coeSubmodule, LieSubmodule.bot_coeSubmodule, + rw [← LieSubmodule.toSubmodule_inj, ker_toSubmodule, LieSubmodule.bot_toSubmodule, LinearMap.ker_eq_bot, coe_toLinearMap] @[simp] -theorem range_coeSubmodule : (f.range : Submodule R L') = LinearMap.range (f : L →ₗ[R] L') := +theorem range_toSubmodule : (f.range : Submodule R L') = LinearMap.range (f : L →ₗ[R] L') := rfl +@[deprecated (since := "2024-12-30")] alias range_coeSubmodule := range_toSubmodule + theorem range_eq_top : f.range = ⊤ ↔ Function.Surjective f := by - rw [← LieSubalgebra.coe_to_submodule_inj, range_coeSubmodule, LieSubalgebra.top_coe_submodule] + rw [← LieSubalgebra.toSubmodule_inj, range_toSubmodule, LieSubalgebra.top_toSubmodule] exact LinearMap.range_eq_top @[simp] theorem idealRange_eq_top_of_surjective (h : Function.Surjective f) : f.idealRange = ⊤ := by rw [← f.range_eq_top] at h - rw [idealRange_eq_lieSpan_range, h, ← LieSubalgebra.coe_to_submodule, ← - LieSubmodule.coe_toSubmodule_inj, LieSubmodule.top_coeSubmodule, - LieSubalgebra.top_coe_submodule, LieSubmodule.coe_lieSpan_submodule_eq_iff] + rw [idealRange_eq_lieSpan_range, h, ← LieSubalgebra.coe_toSubmodule, ← + LieSubmodule.toSubmodule_inj, LieSubmodule.top_toSubmodule, + LieSubalgebra.top_toSubmodule, LieSubmodule.coe_lieSpan_submodule_eq_iff] use ⊤ - exact LieSubmodule.top_coeSubmodule + exact LieSubmodule.top_toSubmodule theorem isIdealMorphism_of_surjective (h : Function.Surjective f) : f.IsIdealMorphism := by rw [isIdealMorphism_def, f.idealRange_eq_top_of_surjective h, f.range_eq_top.mpr h, - LieIdeal.top_coe_lieSubalgebra] + LieIdeal.top_toLieSubalgebra] end LieHom @@ -1120,7 +1191,7 @@ theorem coe_map_of_surjective (h : Function.Surjective f) : obtain ⟨z₂, hz₂, rfl⟩ := hy' obtain ⟨z₁, rfl⟩ := h x simp only [LieHom.coe_toLinearMap, SetLike.mem_coe, Set.mem_image, - LieSubmodule.mem_coeSubmodule, Submodule.mem_carrier, Submodule.map_coe] + LieSubmodule.mem_toSubmodule, Submodule.mem_carrier, Submodule.map_coe] use ⁅z₁, z₂⁆ exact ⟨I.lie_mem hz₂, f.map_lie z₁ z₂⟩ } erw [LieSubmodule.coe_lieSpan_submodule_eq_iff] @@ -1128,7 +1199,7 @@ theorem coe_map_of_surjective (h : Function.Surjective f) : theorem mem_map_of_surjective {y : L'} (h₁ : Function.Surjective f) (h₂ : y ∈ I.map f) : ∃ x : I, f x = y := by - rw [← LieSubmodule.mem_coeSubmodule, coe_map_of_surjective h₁, Submodule.mem_map] at h₂ + rw [← LieSubmodule.mem_toSubmodule, coe_map_of_surjective h₁, Submodule.mem_map] at h₂ obtain ⟨x, hx, rfl⟩ := h₂ use ⟨x, hx⟩ rw [LieHom.coe_toLinearMap] @@ -1179,13 +1250,13 @@ theorem map_comap_eq (h : f.IsIdealMorphism) : map f (comap f J) = f.idealRange apply le_antisymm · rw [le_inf_iff]; exact ⟨f.map_le_idealRange _, map_comap_le⟩ · rw [f.isIdealMorphism_def] at h - rw [← SetLike.coe_subset_coe, LieSubmodule.inf_coe, ← coe_toSubalgebra, h] + rw [← SetLike.coe_subset_coe, LieSubmodule.inf_coe, ← coe_toLieSubalgebra, h] rintro y ⟨⟨x, h₁⟩, h₂⟩; rw [← h₁] at h₂ ⊢; exact mem_map h₂ @[simp] theorem comap_map_eq (h : ↑(map f I) = f '' I) : comap f (map f I) = I ⊔ f.ker := by - rw [← LieSubmodule.coe_toSubmodule_inj, comap_coeSubmodule, I.map_coeSubmodule f h, - LieSubmodule.sup_coe_toSubmodule, f.ker_coeSubmodule, Submodule.comap_map_eq] + rw [← LieSubmodule.toSubmodule_inj, comap_toSubmodule, I.map_toSubmodule f h, + LieSubmodule.sup_toSubmodule, f.ker_toSubmodule, Submodule.comap_map_eq] variable (f I J) @@ -1217,8 +1288,8 @@ theorem ker_incl : I.incl.ker = ⊥ := by ext; simp @[simp] theorem incl_idealRange : I.incl.idealRange = I := by - rw [LieHom.idealRange_eq_lieSpan_range, ← LieSubalgebra.coe_to_submodule, ← - LieSubmodule.coe_toSubmodule_inj, incl_range, coe_to_lieSubalgebra_to_submodule, + rw [LieHom.idealRange_eq_lieSpan_range, ← LieSubalgebra.coe_toSubmodule, ← + LieSubmodule.toSubmodule_inj, incl_range, coe_toLieSubalgebra_toSubmodule, LieSubmodule.coe_lieSpan_submodule_eq_iff] use I @@ -1243,11 +1314,13 @@ def ker : LieSubmodule R L M := LieSubmodule.comap f ⊥ @[simp] -theorem ker_coeSubmodule : (f.ker : Submodule R M) = LinearMap.ker (f : M →ₗ[R] N) := +theorem ker_toSubmodule : (f.ker : Submodule R M) = LinearMap.ker (f : M →ₗ[R] N) := rfl +@[deprecated (since := "2024-12-30")] alias ker_coeSubmodule := ker_toSubmodule + theorem ker_eq_bot : f.ker = ⊥ ↔ Function.Injective f := by - rw [← LieSubmodule.coe_toSubmodule_inj, ker_coeSubmodule, LieSubmodule.bot_coeSubmodule, + rw [← LieSubmodule.toSubmodule_inj, ker_toSubmodule, LieSubmodule.bot_toSubmodule, LinearMap.ker_eq_bot, coe_toLinearMap] variable {f} @@ -1278,9 +1351,11 @@ theorem coe_range : f.range = Set.range f := rfl @[simp] -theorem coeSubmodule_range : f.range = LinearMap.range (f : M →ₗ[R] N) := +theorem toSubmodule_range : f.range = LinearMap.range (f : M →ₗ[R] N) := rfl +@[deprecated (since := "2024-12-30")] alias coeSubmodule_range := toSubmodule_range + @[simp] theorem mem_range (n : N) : n ∈ f.range ↔ ∃ m, f m = n := Iff.rfl @@ -1318,12 +1393,12 @@ theorem ker_incl : N.incl.ker = ⊥ := (LieModuleHom.ker_eq_bot N.incl).mpr <| i @[simp] theorem range_incl : N.incl.range = N := by - simp only [← coe_toSubmodule_inj, LieModuleHom.coeSubmodule_range, incl_coe] + simp only [← toSubmodule_inj, LieModuleHom.toSubmodule_range, incl_coe] rw [Submodule.range_subtype] @[simp] theorem comap_incl_self : comap N.incl N = ⊤ := by - simp only [← coe_toSubmodule_inj, coeSubmodule_comap, incl_coe, top_coeSubmodule] + simp only [← toSubmodule_inj, toSubmodule_comap, incl_coe, top_toSubmodule] rw [Submodule.comap_subtype_self] theorem map_incl_top : (⊤ : LieSubmodule R L N).map N.incl = N := by simp diff --git a/Mathlib/Algebra/Lie/TensorProduct.lean b/Mathlib/Algebra/Lie/TensorProduct.lean index 759d0c87ded6b..e942981a6181a 100644 --- a/Mathlib/Algebra/Lie/TensorProduct.lean +++ b/Mathlib/Algebra/Lie/TensorProduct.lean @@ -108,9 +108,9 @@ theorem coe_liftLie_eq_lift_coe (f : M →ₗ⁅R,L⁆ N →ₗ[R] P) : suffices (liftLie R L M N P f : M ⊗[R] N →ₗ[R] P) = lift R L M N P f by rw [← this, LieModuleHom.coe_toLinearMap] ext m n - simp only [liftLie, LinearEquiv.trans_apply, LieModuleEquiv.coe_to_linearEquiv, - coe_linearMap_maxTrivLinearMapEquivLieModuleHom, coe_maxTrivEquiv_apply, - coe_linearMap_maxTrivLinearMapEquivLieModuleHom_symm] + simp only [liftLie, LinearEquiv.trans_apply, LieModuleEquiv.coe_toLinearEquiv, + toLinearMap_maxTrivLinearMapEquivLieModuleHom, coe_maxTrivEquiv_apply, + toLinearMap_maxTrivLinearMapEquivLieModuleHom_symm] theorem liftLie_apply (f : M →ₗ⁅R,L⁆ N →ₗ[R] P) (m : M) (n : N) : liftLie R L M N P f (m ⊗ₜ n) = f m n := by @@ -132,10 +132,12 @@ nonrec def map (f : M →ₗ⁅R,L⁆ P) (g : N →ₗ⁅R,L⁆ Q) : M ⊗[R] N · intro t₁ t₂ ht₁ ht₂; simp only [ht₁, ht₂, lie_add, LinearMap.map_add] } @[simp] -theorem coe_linearMap_map (f : M →ₗ⁅R,L⁆ P) (g : N →ₗ⁅R,L⁆ Q) : +theorem toLinearMap_map (f : M →ₗ⁅R,L⁆ P) (g : N →ₗ⁅R,L⁆ Q) : (map f g : M ⊗[R] N →ₗ[R] P ⊗[R] Q) = TensorProduct.map (f : M →ₗ[R] P) (g : N →ₗ[R] Q) := rfl +@[deprecated (since := "2024-12-30")] alias coe_linearMap_map := toLinearMap_map + @[simp] nonrec theorem map_tmul (f : M →ₗ⁅R,L⁆ P) (g : N →ₗ⁅R,L⁆ Q) (m : M) (n : N) : map f g (m ⊗ₜ n) = f m ⊗ₜ g n := @@ -196,8 +198,8 @@ applying the action of `L` on `M`, we obtain morphism of Lie modules `f : I ⊗ This lemma states that `⁅I, N⁆ = range f`. -/ theorem lieIdeal_oper_eq_tensor_map_range : ⁅I, N⁆ = ((toModuleHom R L M).comp (mapIncl I N : I ⊗[R] N →ₗ⁅R,L⁆ L ⊗[R] M)).range := by - rw [← coe_toSubmodule_inj, lieIdeal_oper_eq_linear_span, LieModuleHom.coeSubmodule_range, - LieModuleHom.coe_linearMap_comp, LinearMap.range_comp, mapIncl_def, coe_linearMap_map, + rw [← toSubmodule_inj, lieIdeal_oper_eq_linear_span, LieModuleHom.toSubmodule_range, + LieModuleHom.toLinearMap_comp, LinearMap.range_comp, mapIncl_def, toLinearMap_map, TensorProduct.map_range_eq_span_tmul, Submodule.map_span] congr; ext m; constructor · rintro ⟨⟨x, hx⟩, ⟨n, hn⟩, rfl⟩; use x ⊗ₜ n; constructor diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean index ff980b97b8cde..068aae6074787 100644 --- a/Mathlib/Algebra/Lie/TraceForm.lean +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -133,7 +133,7 @@ lemma traceForm_eq_zero_if_mem_lcs_of_mem_ucs {x y : L} (k : ℕ) simp [hy] | succ k ih => rw [LieSubmodule.ucs_succ, LieSubmodule.mem_normalizer] at hy - simp_rw [LieIdeal.lcs_succ, ← LieSubmodule.mem_coeSubmodule, + simp_rw [LieIdeal.lcs_succ, ← LieSubmodule.mem_toSubmodule, LieSubmodule.lieIdeal_oper_eq_linear_span', LieSubmodule.mem_top, true_and] at hx refine Submodule.span_induction ?_ ?_ (fun z w _ _ hz hw ↦ ?_) (fun t z _ hz ↦ ?_) hx · rintro - ⟨z, w, hw, rfl⟩ @@ -189,7 +189,7 @@ lemma trace_toEnd_eq_zero_of_mem_lcs trace R _ (toEnd R L M x) = 0 := by replace hx : x ∈ lowerCentralSeries R L L 1 := antitone_lowerCentralSeries _ _ _ hk hx replace hx : x ∈ Submodule.span R {m | ∃ u v : L, ⁅u, v⁆ = m} := by - rw [lowerCentralSeries_succ, ← LieSubmodule.mem_coeSubmodule, + rw [lowerCentralSeries_succ, ← LieSubmodule.mem_toSubmodule, LieSubmodule.lieIdeal_oper_eq_linear_span'] at hx simpa using hx refine Submodule.span_induction (p := fun x _ ↦ trace R _ (toEnd R L M x) = 0) @@ -225,11 +225,11 @@ lemma traceForm_eq_sum_genWeightSpaceOf fun χ m hm ↦ LieSubmodule.lie_mem _ <| LieSubmodule.lie_mem _ hm 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 _ _) + exact LieSubmodule.toSubmodule_eq_bot (genWeightSpaceOf M _ _) classical - have h := LieSubmodule.iSupIndep_iff_coe_toSubmodule.mp <| iSupIndep_genWeightSpaceOf R L M z + have h := LieSubmodule.iSupIndep_iff_toSubmodule.mp <| iSupIndep_genWeightSpaceOf R L M z have hds := DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top h <| by - simp [← LieSubmodule.iSup_coe_toSubmodule] + simp [← LieSubmodule.iSup_toSubmodule] simp only [LinearMap.coeFn_sum, Finset.sum_apply, traceForm_apply_apply, LinearMap.trace_eq_sum_trace_restrict' hds hfin hxy] exact Finset.sum_congr (by simp) (fun χ _ ↦ rfl) @@ -287,7 +287,7 @@ lemma lowerCentralSeries_one_inf_center_le_ker_traceForm [Module.Free R M] [Modu 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] + LieIdeal.coe_toLieSubalgebra_toSubmodule, LieSubmodule.toSubmodule_eq_bot, h] using lowerCentralSeries_one_inf_center_le_ker_traceForm R L M end LieModule @@ -409,8 +409,8 @@ lemma traceForm_eq_sum_finrank_nsmul_mul (x y : L) : fun χ m hm ↦ LieSubmodule.lie_mem _ <| LieSubmodule.lie_mem _ hm classical have hds := DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top - (LieSubmodule.iSupIndep_iff_coe_toSubmodule.mp <| iSupIndep_genWeightSpace' K L M) - (LieSubmodule.iSup_eq_top_iff_coe_toSubmodule.mp <| iSup_genWeightSpace_eq_top' K L M) + (LieSubmodule.iSupIndep_iff_toSubmodule.mp <| iSupIndep_genWeightSpace' K L M) + (LieSubmodule.iSup_eq_top_iff_toSubmodule.mp <| iSup_genWeightSpace_eq_top' K L M) simp_rw [traceForm_apply_apply, LinearMap.trace_eq_sum_trace_restrict hds hxy, ← traceForm_genWeightSpace_eq K L M _ x y] rfl diff --git a/Mathlib/Algebra/Lie/UniversalEnveloping.lean b/Mathlib/Algebra/Lie/UniversalEnveloping.lean index 514ffe220ed66..262843f473d32 100644 --- a/Mathlib/Algebra/Lie/UniversalEnveloping.lean +++ b/Mathlib/Algebra/Lie/UniversalEnveloping.lean @@ -117,7 +117,7 @@ def lift : (L →ₗ⁅R⁆ A) ≃ (UniversalEnvelopingAlgebra R L →ₐ[R] A) ext -- Porting note: was -- simp only [ι, mkAlgHom, TensorAlgebra.lift_ι_apply, LieHom.coe_toLinearMap, - -- LinearMap.toFun_eq_coe, LinearMap.coe_comp, LieHom.coe_linearMap_comp, + -- LinearMap.toFun_eq_coe, LinearMap.coe_comp, LieHom.toLinearMap_comp, -- AlgHom.comp_toLinearMap, Function.comp_apply, AlgHom.toLinearMap_apply, -- RingQuot.liftAlgHom_mkAlgHom_apply, AlgHom.coe_toLieHom, LieHom.coe_mk] -- extra `rfl` after https://github.com/leanprover/lean4/pull/2644 diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index 4828dcf2ae899..0e9698e5fa0d0 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -306,7 +306,7 @@ theorem zero_genWeightSpace_eq_top_of_nilpotent' [IsNilpotent R L M] : theorem coe_genWeightSpace_of_top (χ : L → R) : (genWeightSpace M (χ ∘ (⊤ : LieSubalgebra R L).incl) : Submodule R M) = genWeightSpace M χ := by ext m - simp only [mem_genWeightSpace, LieSubmodule.mem_coeSubmodule, Subtype.forall] + simp only [mem_genWeightSpace, LieSubmodule.mem_toSubmodule, Subtype.forall] apply forall_congr' simp @@ -478,7 +478,7 @@ lemma posFittingComp_le_iInf_lowerCentralSeries : set F := posFittingComp R L M replace hk : (toEnd R L M x ^ k) m ∈ F := by apply posFittingCompOf_le_posFittingComp R L M x - simp_rw [← LieSubmodule.mem_coeSubmodule, posFittingCompOf, hk k (le_refl k)] + simp_rw [← LieSubmodule.mem_toSubmodule, 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) @@ -589,9 +589,9 @@ variable [IsNoetherian R M] [IsArtinian R M] 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_inj, - LieSubmodule.sup_coe_toSubmodule, LieSubmodule.inf_coe_toSubmodule, - LieSubmodule.top_coeSubmodule, LieSubmodule.bot_coeSubmodule, coe_genWeightSpaceOf_zero] using + simpa only [isCompl_iff, codisjoint_iff, disjoint_iff, ← LieSubmodule.toSubmodule_inj, + LieSubmodule.sup_toSubmodule, LieSubmodule.inf_toSubmodule, + LieSubmodule.top_toSubmodule, LieSubmodule.bot_toSubmodule, 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 @@ -648,7 +648,7 @@ end fitting_decomposition lemma disjoint_genWeightSpaceOf [NoZeroSMulDivisors R M] {x : L} {φ₁ φ₂ : R} (h : φ₁ ≠ φ₂) : Disjoint (genWeightSpaceOf M φ₁ x) (genWeightSpaceOf M φ₂ x) := by - rw [LieSubmodule.disjoint_iff_coe_toSubmodule] + rw [LieSubmodule.disjoint_iff_toSubmodule] dsimp [genWeightSpaceOf] exact Module.End.disjoint_genEigenspace _ h _ _ @@ -669,8 +669,8 @@ lemma injOn_genWeightSpace [NoZeroSMulDivisors R M] : See also `LieModule.iSupIndep_genWeightSpace'`. -/ lemma iSupIndep_genWeightSpace [NoZeroSMulDivisors R M] : iSupIndep fun χ : L → R ↦ genWeightSpace M χ := by - simp only [LieSubmodule.iSupIndep_iff_coe_toSubmodule, genWeightSpace, - LieSubmodule.iInf_coe_toSubmodule] + simp only [LieSubmodule.iSupIndep_iff_toSubmodule, genWeightSpace, + LieSubmodule.iInf_toSubmodule] exact Module.End.independent_iInf_maxGenEigenspace_of_forall_mapsTo (toEnd R L M) (fun x y φ z ↦ (genWeightSpaceOf M φ y).lie_mem) @@ -685,7 +685,7 @@ lemma iSupIndep_genWeightSpace' [NoZeroSMulDivisors R M] : lemma iSupIndep_genWeightSpaceOf [NoZeroSMulDivisors R M] (x : L) : iSupIndep fun (χ : R) ↦ genWeightSpaceOf M χ x := by - rw [LieSubmodule.iSupIndep_iff_coe_toSubmodule] + rw [LieSubmodule.iSupIndep_iff_toSubmodule] dsimp [genWeightSpaceOf] exact (toEnd R L M x).independent_genEigenspace _ @@ -728,8 +728,8 @@ instance [IsTriangularizable R L M] : IsTriangularizable R (LieModule.toEnd R L @[simp] lemma iSup_genWeightSpaceOf_eq_top [IsTriangularizable R L M] (x : L) : ⨆ (φ : R), genWeightSpaceOf M φ x = ⊤ := by - rw [← LieSubmodule.coe_toSubmodule_inj, LieSubmodule.iSup_coe_toSubmodule, - LieSubmodule.top_coeSubmodule] + rw [← LieSubmodule.toSubmodule_inj, LieSubmodule.iSup_toSubmodule, + LieSubmodule.top_toSubmodule] dsimp [genWeightSpaceOf] exact IsTriangularizable.maxGenEigenspace_eq_top x @@ -766,8 +766,8 @@ instance (N : LieSubmodule K L M) [IsTriangularizable K L M] : IsTriangularizabl See also `LieModule.iSup_genWeightSpace_eq_top'`. -/ lemma iSup_genWeightSpace_eq_top [IsTriangularizable K L M] : ⨆ χ : L → K, genWeightSpace M χ = ⊤ := by - simp only [← LieSubmodule.coe_toSubmodule_inj, LieSubmodule.iSup_coe_toSubmodule, - LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.top_coeSubmodule, genWeightSpace] + simp only [← LieSubmodule.toSubmodule_inj, LieSubmodule.iSup_toSubmodule, + LieSubmodule.iInf_toSubmodule, LieSubmodule.top_toSubmodule, genWeightSpace] refine Module.End.iSup_iInf_maxGenEigenspace_eq_top_of_forall_mapsTo (toEnd K L M) (fun x y φ z ↦ (genWeightSpaceOf M φ y).lie_mem) ?_ apply IsTriangularizable.maxGenEigenspace_eq_top diff --git a/Mathlib/Algebra/Lie/Weights/Cartan.lean b/Mathlib/Algebra/Lie/Weights/Cartan.lean index 576a1918640c9..17a743e544aa1 100644 --- a/Mathlib/Algebra/Lie/Weights/Cartan.lean +++ b/Mathlib/Algebra/Lie/Weights/Cartan.lean @@ -195,8 +195,8 @@ instance [Nontrivial H] : Nontrivial (genWeightSpace L (0 : H → R)) := by ⟨y, toLieSubmodule_le_rootSpace_zero R L H hy⟩, by simpa using e⟩ theorem le_zeroRootSubalgebra : H ≤ zeroRootSubalgebra R L H := by - rw [← LieSubalgebra.coe_submodule_le_coe_submodule, ← H.coe_toLieSubmodule, - coe_zeroRootSubalgebra, LieSubmodule.coeSubmodule_le_coeSubmodule] + rw [← LieSubalgebra.toSubmodule_le_toSubmodule, ← H.coe_toLieSubmodule, + coe_zeroRootSubalgebra, LieSubmodule.toSubmodule_le_toSubmodule] exact toLieSubmodule_le_rootSpace_zero R L H @[simp] @@ -240,7 +240,7 @@ theorem zeroRootSubalgebra_eq_iff_is_cartan [IsNoetherian R L] : @[simp] theorem rootSpace_zero_eq (H : LieSubalgebra R L) [H.IsCartanSubalgebra] [IsNoetherian R L] : rootSpace H 0 = H.toLieSubmodule := by - rw [← LieSubmodule.coe_toSubmodule_inj, ← coe_zeroRootSubalgebra, + rw [← LieSubmodule.toSubmodule_inj, ← coe_zeroRootSubalgebra, zeroRootSubalgebra_eq_of_is_cartan R L H, LieSubalgebra.coe_toLieSubmodule] variable {R L H} @@ -271,8 +271,8 @@ lemma mem_corootSpace {x : H} : simp only [rootSpaceProduct_def, LieModuleHom.mem_range, LieSubmodule.mem_map, LieSubmodule.incl_apply, SetLike.coe_eq_coe, exists_eq_right] rfl - simp_rw [this, corootSpace, ← LieModuleHom.map_top, ← LieSubmodule.mem_coeSubmodule, - LieSubmodule.coeSubmodule_map, LieSubmodule.top_coeSubmodule, ← TensorProduct.span_tmul_eq_top, + simp_rw [this, corootSpace, ← LieModuleHom.map_top, ← LieSubmodule.mem_toSubmodule, + LieSubmodule.toSubmodule_map, LieSubmodule.top_toSubmodule, ← TensorProduct.span_tmul_eq_top, LinearMap.map_span, Set.image, Set.mem_setOf_eq, exists_exists_exists_and_eq] change (x : L) ∈ Submodule.span R {x | ∃ (a : rootSpace H α) (b : rootSpace H (-α)), ⁅(a : L), (b : L)⁆ = x} ↔ _ @@ -289,7 +289,7 @@ lemma mem_corootSpace' {x : H} : simp_rw [SetLike.mem_coe] rw [← Submodule.mem_map, Submodule.coe_subtype, Submodule.map_span, mem_corootSpace, ← this] ext u - simp only [Submodule.coe_subtype, mem_image, Subtype.exists, LieSubalgebra.mem_coe_submodule, + simp only [Submodule.coe_subtype, mem_image, Subtype.exists, LieSubalgebra.mem_toSubmodule, 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⟩⟩ diff --git a/Mathlib/Algebra/Lie/Weights/Chain.lean b/Mathlib/Algebra/Lie/Weights/Chain.lean index cc0efd7803997..b2ba09bb4e2bd 100644 --- a/Mathlib/Algebra/Lie/Weights/Chain.lean +++ b/Mathlib/Algebra/Lie/Weights/Chain.lean @@ -208,12 +208,12 @@ lemma exists_forall_mem_corootSpace_smul_add_eq_zero refine ⟨a, b, Int.ofNat_pos.mpr hb, fun x hx ↦ ?_⟩ let N : ℤ → Submodule R M := fun k ↦ genWeightSpace M (k • α + χ) have h₁ : iSupIndep fun (i : Finset.Ioo p q) ↦ N i := by - rw [← LieSubmodule.iSupIndep_iff_coe_toSubmodule] + rw [← LieSubmodule.iSupIndep_iff_toSubmodule] refine (iSupIndep_genWeightSpace R H M).comp fun i j hij ↦ ?_ exact SetCoe.ext <| smul_left_injective ℤ hα <| by rwa [add_left_inj] at hij have h₂ : ∀ i, MapsTo (toEnd R H M x) ↑(N i) ↑(N i) := fun _ _ ↦ LieSubmodule.lie_mem _ have h₃ : genWeightSpaceChain M α χ p q = ⨆ i ∈ Finset.Ioo p q, N i := by - simp_rw [N, genWeightSpaceChain_def', LieSubmodule.iSup_coe_toSubmodule] + simp_rw [N, genWeightSpaceChain_def', LieSubmodule.iSup_toSubmodule] rw [← trace_toEnd_genWeightSpaceChain_eq_zero M α χ p q hp hq hx, ← LieSubmodule.toEnd_restrict_eq_toEnd] -- The lines below illustrate the cost of treating `LieSubmodule` as both a diff --git a/Mathlib/Algebra/Lie/Weights/Killing.lean b/Mathlib/Algebra/Lie/Weights/Killing.lean index fa7a40102df4a..48d999e1a7314 100644 --- a/Mathlib/Algebra/Lie/Weights/Killing.lean +++ b/Mathlib/Algebra/Lie/Weights/Killing.lean @@ -54,8 +54,8 @@ lemma ker_restrict_eq_bot_of_isCartanSubalgebra have h : Codisjoint (rootSpace H 0) (LieModule.posFittingComp R H L) := (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_inj, LieSubmodule.sup_coe_toSubmodule, - LieSubmodule.top_coeSubmodule, rootSpace_zero_eq R L H, LieSubalgebra.coe_toLieSubmodule, + rwa [codisjoint_iff, ← LieSubmodule.toSubmodule_inj, LieSubmodule.sup_toSubmodule, + LieSubmodule.top_toSubmodule, rootSpace_zero_eq R L H, LieSubalgebra.coe_toLieSubmodule, ← codisjoint_iff] at h suffices this : ∀ m₀ ∈ H, ∀ m₁ ∈ LieModule.posFittingComp R H L, killingForm R L m₀ m₁ = 0 by simp [LinearMap.BilinForm.ker_restrict_eq_of_codisjoint h this] @@ -113,8 +113,8 @@ lemma killingForm_apply_eq_zero_of_mem_rootSpace_of_add_ne_zero {α β : H → K mapsTo_toEnd_genWeightSpace_add_of_mem_rootSpace K L H L β γ hy classical have hds := DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top - (LieSubmodule.iSupIndep_iff_coe_toSubmodule.mp <| iSupIndep_genWeightSpace K H L) - (LieSubmodule.iSup_eq_top_iff_coe_toSubmodule.mp <| iSup_genWeightSpace_eq_top K H L) + (LieSubmodule.iSupIndep_iff_toSubmodule.mp <| iSupIndep_genWeightSpace K H L) + (LieSubmodule.iSup_eq_top_iff_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 @@ -337,7 +337,7 @@ lemma coe_corootSpace_eq_span_singleton' (α : Weight K H L) : rw [lie_eq_killingForm_smul_of_mem_rootSpace_of_mem_rootSpace_neg heα hfα, SetLike.mem_coe, Submodule.mem_span_singleton] exact ⟨killingForm K L e f, rfl⟩ - simp only [LieSubmodule.mem_coeSubmodule, mem_corootSpace] at hx' + simp only [LieSubmodule.mem_toSubmodule, mem_corootSpace] at hx' replace this := Submodule.span_mono this hx' rw [Submodule.span_span] at this rw [Submodule.mem_span_singleton] at this ⊢ @@ -345,7 +345,7 @@ lemma coe_corootSpace_eq_span_singleton' (α : Weight K H L) : use t simp only [Subtype.ext_iff] rw [Submodule.coe_smul_of_tower] - · simp only [Submodule.span_singleton_le_iff_mem, LieSubmodule.mem_coeSubmodule] + · simp only [Submodule.span_singleton_le_iff_mem, LieSubmodule.mem_toSubmodule] exact cartanEquivDual_symm_apply_mem_corootSpace α end PerfectField @@ -412,7 +412,7 @@ lemma coe_corootSpace_eq_span_singleton (α : Weight K H L) : @[simp] lemma corootSpace_eq_bot_iff {α : Weight K H L} : corootSpace α = ⊥ ↔ α.IsZero := by - simp [← LieSubmodule.coeSubmodule_eq_bot_iff, coe_corootSpace_eq_span_singleton α] + simp [← LieSubmodule.toSubmodule_eq_bot, coe_corootSpace_eq_span_singleton α] lemma isCompl_ker_weight_span_coroot (α : Weight K H L) : IsCompl α.ker (K ∙ coroot α) := by @@ -424,12 +424,12 @@ lemma isCompl_ker_weight_span_coroot (α : Weight K H L) : apply Module.Dual.isCompl_ker_of_disjoint_of_ne_bot (by aesop) (disjoint_ker_weight_corootSpace α) replace hα : corootSpace α ≠ ⊥ := by simpa using hα - rwa [ne_eq, ← LieSubmodule.coe_toSubmodule_inj] at hα + rwa [ne_eq, ← LieSubmodule.toSubmodule_inj] at hα lemma traceForm_eq_zero_of_mem_ker_of_mem_span_coroot {α : Weight K H L} {x y : H} (hx : x ∈ α.ker) (hy : y ∈ K ∙ coroot α) : traceForm K H L x y = 0 := by - rw [← coe_corootSpace_eq_span_singleton, LieSubmodule.mem_coeSubmodule, mem_corootSpace'] at hy + rw [← coe_corootSpace_eq_span_singleton, LieSubmodule.mem_toSubmodule, mem_corootSpace'] at hy induction hy using Submodule.span_induction with | mem z hz => obtain ⟨u, hu, v, -, huv⟩ := hz diff --git a/Mathlib/Algebra/Lie/Weights/Linear.lean b/Mathlib/Algebra/Lie/Weights/Linear.lean index c2ca4eedbf3b4..596b044234646 100644 --- a/Mathlib/Algebra/Lie/Weights/Linear.lean +++ b/Mathlib/Algebra/Lie/Weights/Linear.lean @@ -97,14 +97,14 @@ instance instLinearWeightsOfIsLieAbelian [IsLieAbelian L] [NoZeroSMulDivisors R 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_inj, genWeightSpace, genWeightSpaceOf, - LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.bot_coeSubmodule] at hχ + simp_rw [Ne, ← LieSubmodule.toSubmodule_inj, genWeightSpace, genWeightSpaceOf, + LieSubmodule.iInf_toSubmodule, LieSubmodule.bot_toSubmodule] 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_inj, genWeightSpace, genWeightSpaceOf, - LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.bot_coeSubmodule] at hχ + simp_rw [Ne, ← LieSubmodule.toSubmodule_inj, genWeightSpace, genWeightSpaceOf, + LieSubmodule.iInf_toSubmodule, LieSubmodule.bot_toSubmodule] at hχ exact Module.End.map_smul_of_iInf_genEigenspace_ne_bot (toEnd R L M).toLinearMap χ _ hχ t x map_lie := fun χ hχ t x ↦ by From d458ec9af73e637e20410d0496462c2426f887c3 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Thu, 2 Jan 2025 11:03:51 +0000 Subject: [PATCH 077/189] feat(NumberTheory/LSeries/Injectivity): new file (#20370) This adds the statement that a converging L-series determines its coefficients. --- Mathlib.lean | 1 + .../SpecialFunctions/Pow/Complex.lean | 3 + Mathlib/NumberTheory/LSeries/Basic.lean | 5 + Mathlib/NumberTheory/LSeries/Convergence.lean | 14 + Mathlib/NumberTheory/LSeries/Convolution.lean | 6 + Mathlib/NumberTheory/LSeries/Injectivity.lean | 241 ++++++++++++++++++ 6 files changed, 270 insertions(+) create mode 100644 Mathlib/NumberTheory/LSeries/Injectivity.lean diff --git a/Mathlib.lean b/Mathlib.lean index 6ff4958551d47..c07b91155b702 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3908,6 +3908,7 @@ import Mathlib.NumberTheory.LSeries.HurwitzZeta import Mathlib.NumberTheory.LSeries.HurwitzZetaEven import Mathlib.NumberTheory.LSeries.HurwitzZetaOdd import Mathlib.NumberTheory.LSeries.HurwitzZetaValues +import Mathlib.NumberTheory.LSeries.Injectivity import Mathlib.NumberTheory.LSeries.Linearity import Mathlib.NumberTheory.LSeries.MellinEqDirichlet import Mathlib.NumberTheory.LSeries.Nonvanishing diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean index 1afcd23913d7d..2d2f3a84e0de5 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean @@ -234,6 +234,9 @@ theorem conj_cpow (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : conj x ^ n = conj (x theorem cpow_conj (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : x ^ conj n = conj (conj x ^ n) := by rw [conj_cpow _ _ hx, conj_conj] +lemma cpow_natCast_add_one_ne_zero (n : ℕ) (z : ℂ) : (n + 1 : ℂ) ^ z ≠ 0 := + mt (cpow_eq_zero_iff ..).mp fun H ↦ by norm_cast at H; exact H.1 + end Complex -- section Tactics diff --git a/Mathlib/NumberTheory/LSeries/Basic.lean b/Mathlib/NumberTheory/LSeries/Basic.lean index 99720f95d8532..b6186de00d049 100644 --- a/Mathlib/NumberTheory/LSeries/Basic.lean +++ b/Mathlib/NumberTheory/LSeries/Basic.lean @@ -91,6 +91,10 @@ lemma term_congr {f g : ℕ → ℂ} (h : ∀ {n}, n ≠ 0 → f n = g n) (s : term f s n = term g s n := by rcases eq_or_ne n 0 with hn | hn <;> simp [hn, h] +lemma pow_mul_term_eq (f : ℕ → ℂ) (s : ℂ) (n : ℕ) : + (n + 1) ^ s * term f s (n + 1) = f (n + 1) := by + simp [term, cpow_natCast_add_one_ne_zero n _, mul_comm (f _), mul_div_assoc'] + lemma norm_term_eq (f : ℕ → ℂ) (s : ℂ) (n : ℕ) : ‖term f s n‖ = if n = 0 then 0 else ‖f n‖ / n ^ s.re := by rcases eq_or_ne n 0 with rfl | hn @@ -144,6 +148,7 @@ noncomputable def LSeries (f : ℕ → ℂ) (s : ℂ) : ℂ := ∑' n, term f s n +-- TODO: change argument order in `LSeries_congr` to have `s` last. lemma LSeries_congr {f g : ℕ → ℂ} (s : ℂ) (h : ∀ {n}, n ≠ 0 → f n = g n) : LSeries f s = LSeries g s := tsum_congr <| term_congr h s diff --git a/Mathlib/NumberTheory/LSeries/Convergence.lean b/Mathlib/NumberTheory/LSeries/Convergence.lean index 26ca9bd9f6e91..d54f658f77e00 100644 --- a/Mathlib/NumberTheory/LSeries/Convergence.lean +++ b/Mathlib/NumberTheory/LSeries/Convergence.lean @@ -137,3 +137,17 @@ lemma LSeries.summable_real_of_abscissaOfAbsConv_lt {f : ℕ → ℝ} {x : ℝ} filter_upwards [Set.Finite.compl_mem_cofinite <| Set.finite_singleton 0] with n hn simp only [Set.mem_compl_iff, Set.mem_singleton_iff] at hn exact if_neg hn + +/-- If `F` is a binary operation on `ℕ → ℂ` with the property that the `LSeries` of `F f g` +converges whenever the `LSeries` of `f` and `g` do, then the abscissa of absolute convergence +of `F f g` is at most the maximum of the abscissa of absolute convergence of `f` +and that of `g`. -/ +lemma LSeries.abscissaOfAbsConv_binop_le {F : (ℕ → ℂ) → (ℕ → ℂ) → (ℕ → ℂ)} + (hF : ∀ {f g s}, LSeriesSummable f s → LSeriesSummable g s → LSeriesSummable (F f g) s) + (f g : ℕ → ℂ) : + abscissaOfAbsConv (F f g) ≤ max (abscissaOfAbsConv f) (abscissaOfAbsConv g) := by + refine abscissaOfAbsConv_le_of_forall_lt_LSeriesSummable' fun x hx ↦ hF ?_ ?_ + · exact LSeriesSummable_of_abscissaOfAbsConv_lt_re <| + (ofReal_re x).symm ▸ (le_max_left ..).trans_lt hx + · exact LSeriesSummable_of_abscissaOfAbsConv_lt_re <| + (ofReal_re x).symm ▸ (le_max_right ..).trans_lt hx diff --git a/Mathlib/NumberTheory/LSeries/Convolution.lean b/Mathlib/NumberTheory/LSeries/Convolution.lean index 5063a90643a41..eed4c45fbce01 100644 --- a/Mathlib/NumberTheory/LSeries/Convolution.lean +++ b/Mathlib/NumberTheory/LSeries/Convolution.lean @@ -167,6 +167,12 @@ lemma LSeriesSummable.convolution {f g : ℕ → ℂ} {s : ℂ} (hf : LSeriesSum LSeriesSummable (f ⍟ g) s := (LSeriesHasSum.convolution hf.LSeriesHasSum hg.LSeriesHasSum).LSeriesSummable +/-- The abscissa of absolute convergence of `f ⍟ g` is at most the maximum of those +of `f` and `g`. -/ +lemma LSeries.abscissaOfAbsConv_convolution_le (f g : ℕ → ℂ) : + abscissaOfAbsConv (f ⍟ g) ≤ max (abscissaOfAbsConv f) (abscissaOfAbsConv g) := + abscissaOfAbsConv_binop_le LSeriesSummable.convolution f g + namespace ArithmeticFunction /-! diff --git a/Mathlib/NumberTheory/LSeries/Injectivity.lean b/Mathlib/NumberTheory/LSeries/Injectivity.lean new file mode 100644 index 0000000000000..aead49d62c1c3 --- /dev/null +++ b/Mathlib/NumberTheory/LSeries/Injectivity.lean @@ -0,0 +1,241 @@ +/- +Copyright (c) 2024 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Stoll +-/ +import Mathlib.Analysis.InnerProductSpace.Basic +import Mathlib.Analysis.Normed.Group.Tannery +import Mathlib.NumberTheory.LSeries.Convergence +import Mathlib.NumberTheory.LSeries.Linearity + +/-! +# A converging L-series determines its coefficients + +We show that two functions `f` and `g : ℕ → ℂ` whose L-series agree and both converge somewhere +must agree on all nonzero arguments. See `LSeries_eq_iff_of_abscissaOfAbsConv_lt_top` +and `LSeries_injOn`. +-/ + +open LSeries Complex + +-- The following two lemmas need both `LSeries.Linearity` and `LSeries.Convergence`, +-- so cannot live in either of these files. + +/-- The abscissa of absolute convergence of `f + g` is at most the maximum of those +of `f` and `g`. -/ +lemma LSeries.abscissaOfAbsConv_add_le (f g : ℕ → ℂ) : + abscissaOfAbsConv (f + g) ≤ max (abscissaOfAbsConv f) (abscissaOfAbsConv g) := + abscissaOfAbsConv_binop_le LSeriesSummable.add f g + +/-- The abscissa of absolute convergence of `f - g` is at most the maximum of those +of `f` and `g`. -/ +lemma LSeries.abscissaOfAbsConv_sub_le (f g : ℕ → ℂ) : + abscissaOfAbsConv (f - g) ≤ max (abscissaOfAbsConv f) (abscissaOfAbsConv g) := + abscissaOfAbsConv_binop_le LSeriesSummable.sub f g + +private +lemma cpow_mul_div_cpow_eq_div_div_cpow (m n : ℕ) (z : ℂ) (x : ℝ) : + (n + 1) ^ (x : ℂ) * (z / m ^ (x : ℂ)) = z / (m / (n + 1)) ^ (x : ℂ) := by + have Hn : (0 : ℝ) ≤ (n + 1 : ℝ)⁻¹ := by positivity + rw [← mul_div_assoc, mul_comm, div_eq_mul_inv z, mul_div_assoc] + congr + simp_rw [div_eq_mul_inv] + rw [show (n + 1 : ℂ)⁻¹ = (n + 1 : ℝ)⁻¹ by simp, + show (n + 1 : ℂ) = (n + 1 : ℝ) by norm_cast, show (m : ℂ) = (m : ℝ) by norm_cast, + mul_cpow_ofReal_nonneg m.cast_nonneg Hn, mul_inv, mul_comm] + congr + rw [← cpow_neg, show (-x : ℂ) = (-1 : ℝ) * x by simp, cpow_mul_ofReal_nonneg Hn, + Real.rpow_neg_one, inv_inv] + +open Filter Real in +/-- If the coefficients `f m` of an L-series are zero for `m ≤ n` and the L-series converges +at some point, then `f (n+1)` is the limit of `(n+1)^x * LSeries f x` as `x → ∞`. -/ +lemma LSeries.tendsto_cpow_mul_atTop {f : ℕ → ℂ} {n : ℕ} (h : ∀ m ≤ n, f m = 0) + (ha : abscissaOfAbsConv f < ⊤): + Tendsto (fun x : ℝ ↦ (n + 1) ^ (x : ℂ) * LSeries f x) atTop (nhds (f (n + 1))) := by + obtain ⟨y, hay, hyt⟩ := exists_between ha + lift y to ℝ using ⟨hyt.ne, ((OrderBot.bot_le _).trans_lt hay).ne'⟩ + -- `F x m` is the `m`th term of `(n+1)^x * LSeries f x`, except that `F x (n+1) = 0` + let F := fun (x : ℝ) ↦ {m | n + 1 < m}.indicator (fun m ↦ f m / (m / (n + 1) : ℂ) ^ (x : ℂ)) + have hF₀ (x : ℝ) {m : ℕ} (hm : m ≤ n + 1) : F x m = 0 := by simp [F, not_lt_of_le hm] + have hF (x : ℝ) {m : ℕ} (hm : m ≠ n + 1) : F x m = ((n + 1) ^ (x : ℂ)) * term f x m := by + rcases lt_trichotomy m (n + 1) with H | rfl | H + · simp [Nat.not_lt_of_gt H, term, h m <| Nat.lt_succ_iff.mp H, F] + · exact (hm rfl).elim + · simp [H, term, (n.zero_lt_succ.trans H).ne', F, cpow_mul_div_cpow_eq_div_div_cpow] + have hs {x : ℝ} (hx : x ≥ y) : Summable fun m ↦ (n + 1) ^ (x : ℂ) * term f x m := by + refine (summable_mul_left_iff <| cpow_natCast_add_one_ne_zero n _).mpr <| + LSeriesSummable_of_abscissaOfAbsConv_lt_re ?_ + simpa only [ofReal_re] using hay.trans_le <| EReal.coe_le_coe_iff.mpr hx + -- we can write `(n+1)^x * LSeries f x` as `f (n+1)` plus the series over `F x` + have key : ∀ x ≥ y, (n + 1) ^ (x : ℂ) * LSeries f x = f (n + 1) + ∑' m : ℕ, F x m := by + intro x hx + rw [LSeries, ← tsum_mul_left, tsum_eq_add_tsum_ite (hs hx) (n + 1), pow_mul_term_eq f x n] + congr + ext1 m + rcases eq_or_ne m (n + 1) with rfl | hm + · simp [hF₀ x le_rfl] + · simp [hm, hF] + -- reduce to showing that `∑' m, F x m → 0` as `x → ∞` + conv => enter [3, 1]; rw [← add_zero (f _)] + refine Tendsto.congr' + (eventuallyEq_of_mem (s := {x | y ≤ x}) (mem_atTop y) key).symm <| tendsto_const_nhds.add ?_ + -- get the prerequisites for applying dominated convergence + have hys : Summable (F y) := by + refine ((hs le_rfl).indicator {m | n + 1 < m}).congr fun m ↦ ?_ + by_cases hm : n + 1 < m + · simp [hF, hm, hm.ne'] + · simp [hm, hF₀ _ (le_of_not_lt hm)] + have hc (k : ℕ) : Tendsto (F · k) atTop (nhds 0) := by + rcases lt_or_le (n + 1) k with H | H + · have H₀ : (0 : ℝ) ≤ k / (n + 1) := by positivity + have H₀' : (0 : ℝ) ≤ (n + 1) / k := by positivity + have H₁ : (k / (n + 1) : ℂ) = (k / (n + 1) : ℝ) := by push_cast; rfl + have H₂ : (n + 1) / k < (1 : ℝ) := + (div_lt_one <| mod_cast n.succ_pos.trans H).mpr <| mod_cast H + simp only [Set.mem_setOf_eq, H, Set.indicator_of_mem, F] + conv => + enter [1, x] + rw [div_eq_mul_inv, H₁, ← ofReal_cpow H₀, ← ofReal_inv, ← Real.inv_rpow H₀, inv_div] + conv => enter [3, 1]; rw [← mul_zero (f k)] + exact + (tendsto_rpow_atTop_of_base_lt_one _ (neg_one_lt_zero.trans_le H₀') H₂).ofReal.const_mul _ + · simp [hF₀ _ H] + rw [show (0 : ℂ) = tsum (fun _ : ℕ ↦ 0) from tsum_zero.symm] + refine tendsto_tsum_of_dominated_convergence hys.norm hc <| eventually_iff.mpr ?_ + filter_upwards [mem_atTop y] with y' hy' k + -- it remains to show that `‖F y' k‖ ≤ ‖F y k‖` (for `y' ≥ y`) + rcases lt_or_le (n + 1) k with H | H + · simp only [Set.mem_setOf_eq, H, Set.indicator_of_mem, norm_div, Complex.norm_eq_abs, + abs_cpow_real, map_div₀, abs_natCast, F] + rw [← Nat.cast_one, ← Nat.cast_add, abs_natCast] + have hkn : 1 ≤ (k / (n + 1 :) : ℝ) := + (one_le_div (by positivity)).mpr <| mod_cast Nat.le_of_succ_le H + exact div_le_div_of_nonneg_left (Complex.abs.nonneg _) + (rpow_pos_of_pos (zero_lt_one.trans_le hkn) _) <| rpow_le_rpow_of_exponent_le hkn hy' + · simp [hF₀ _ H] + +open Filter in +/-- If the L-series of `f` converges at some point, then `f 1` is the limit of `LSeries f x` +as `x → ∞`. -/ +lemma LSeries.tendsto_atTop {f : ℕ → ℂ} (ha : abscissaOfAbsConv f < ⊤): + Tendsto (fun x : ℝ ↦ LSeries f x) atTop (nhds (f 1)) := by + let F (n : ℕ) : ℂ := if n = 0 then 0 else f n + have hF₀ : F 0 = 0 := rfl + have hF {n : ℕ} (hn : n ≠ 0) : F n = f n := if_neg hn + have ha' : abscissaOfAbsConv F < ⊤ := (abscissaOfAbsConv_congr hF).symm ▸ ha + simp_rw [← LSeries_congr _ hF] + convert LSeries.tendsto_cpow_mul_atTop (n := 0) (fun _ hm ↦ Nat.le_zero.mp hm ▸ hF₀) ha' using 1 + simp + +lemma LSeries_eq_zero_of_abscissaOfAbsConv_eq_top {f : ℕ → ℂ} (h : abscissaOfAbsConv f = ⊤) : + LSeries f = 0 := by + ext1 s + exact LSeries.eq_zero_of_not_LSeriesSummable f s <| mt LSeriesSummable.abscissaOfAbsConv_le <| + h ▸ fun H ↦ (H.trans_lt <| EReal.coe_lt_top _).false + +open Filter Nat in +/-- The `LSeries` of `f` is zero for large real arguments if and only if either `f n = 0` +for all `n ≠ 0` or the L-series converges nowhere. -/ +lemma LSeries_eventually_eq_zero_iff' {f : ℕ → ℂ} : + (fun x : ℝ ↦ LSeries f x) =ᶠ[atTop] 0 ↔ (∀ n ≠ 0, f n = 0) ∨ abscissaOfAbsConv f = ⊤ := by + by_cases h : abscissaOfAbsConv f = ⊤ <;> simp [h] + · exact Eventually.of_forall <| by simp [LSeries_eq_zero_of_abscissaOfAbsConv_eq_top h] + · refine ⟨fun H ↦ ?_, fun H ↦ Eventually.of_forall fun x ↦ ?_⟩ + · let F (n : ℕ) : ℂ := if n = 0 then 0 else f n + have hF₀ : F 0 = 0 := rfl + have hF {n : ℕ} (hn : n ≠ 0) : F n = f n := if_neg hn + suffices ∀ n, F n = 0 from fun n hn ↦ (hF hn).symm.trans (this n) + have ha : ¬ abscissaOfAbsConv F = ⊤ := abscissaOfAbsConv_congr hF ▸ h + have h' (x : ℝ) : LSeries F x = LSeries f x := LSeries_congr x hF + have H' (n : ℕ) : (fun x : ℝ ↦ n ^ (x : ℂ) * LSeries F x) =ᶠ[atTop] fun _ ↦ 0 := by + simp only [h'] + rw [eventuallyEq_iff_exists_mem] at H ⊢ + obtain ⟨s, hs⟩ := H + exact ⟨s, hs.1, fun x hx ↦ by simp [hs.2 hx]⟩ + intro n + induction' n using Nat.strongRecOn with n ih + -- it suffices to show that `n ^ x * LSeries F x` tends to `F n` as `x` tends to `∞` + suffices Tendsto (fun x : ℝ ↦ n ^ (x : ℂ) * LSeries F x) atTop (nhds (F n)) by + replace this := this.congr' <| H' n + simp only [tendsto_const_nhds_iff] at this + exact this.symm + cases n with + | zero => exact Tendsto.congr' (H' 0).symm <| by simp [hF₀] + | succ n => + simpa using LSeries.tendsto_cpow_mul_atTop (fun m hm ↦ ih m <| lt_succ_of_le hm) <| + Ne.lt_top ha + · simp [LSeries_congr x fun {n} ↦ H n, show (fun _ : ℕ ↦ (0 : ℂ)) = 0 from rfl] + +open Nat in +/-- Assuming `f 0 = 0`, the `LSeries` of `f` is zero if and only if either `f = 0` or the +L-series converges nowhere. -/ +lemma LSeries_eq_zero_iff {f : ℕ → ℂ} (hf : f 0 = 0) : + LSeries f = 0 ↔ f = 0 ∨ abscissaOfAbsConv f = ⊤ := by + by_cases h : abscissaOfAbsConv f = ⊤ <;> simp [h] + · exact LSeries_eq_zero_of_abscissaOfAbsConv_eq_top h + · refine ⟨fun H ↦ ?_, fun H ↦ H ▸ LSeries_zero⟩ + convert (LSeries_eventually_eq_zero_iff'.mp ?_).resolve_right h + · refine ⟨fun H' _ _ ↦ by rw [H', Pi.zero_apply], fun H' ↦ ?_⟩ + ext (- | m) + · simp [hf] + · simp [H'] + · simpa only [H] using Filter.EventuallyEq.rfl + +open Filter in +/-- If the `LSeries` of `f` and of `g` converge somewhere and agree on large real arguments, +then the L-series of `f - g` is zero for large real arguments. -/ +lemma LSeries_sub_eventuallyEq_zero_of_LSeries_eventually_eq {f g : ℕ → ℂ} + (hf : abscissaOfAbsConv f < ⊤) (hg : abscissaOfAbsConv g < ⊤) + (h : (fun x : ℝ ↦ LSeries f x) =ᶠ[atTop] fun x ↦ LSeries g x) : + (fun x : ℝ ↦ LSeries (f - g) x) =ᶠ[atTop] (0 : ℝ → ℂ) := by + rw [EventuallyEq, eventually_atTop] at h ⊢ + obtain ⟨x₀, hx₀⟩ := h + obtain ⟨yf, hyf₁, hyf₂⟩ := exists_between hf + obtain ⟨yg, hyg₁, hyg₂⟩ := exists_between hg + lift yf to ℝ using ⟨hyf₂.ne, ((OrderBot.bot_le _).trans_lt hyf₁).ne'⟩ + lift yg to ℝ using ⟨hyg₂.ne, ((OrderBot.bot_le _).trans_lt hyg₁).ne'⟩ + refine ⟨max x₀ (max yf yg), fun x hx ↦ ?_⟩ + have Hf : LSeriesSummable f x := by + refine LSeriesSummable_of_abscissaOfAbsConv_lt_re <| + (ofReal_re x).symm ▸ hyf₁.trans_le (EReal.coe_le_coe_iff.mpr ?_) + exact (le_max_left _ yg).trans <| (le_max_right x₀ _).trans hx + have Hg : LSeriesSummable g x := by + refine LSeriesSummable_of_abscissaOfAbsConv_lt_re <| + (ofReal_re x).symm ▸ hyg₁.trans_le (EReal.coe_le_coe_iff.mpr ?_) + exact (le_max_right yf _).trans <| (le_max_right x₀ _).trans hx + rw [LSeries_sub Hf Hg, hx₀ x <| (le_max_left ..).trans hx, sub_self, Pi.zero_apply] + +open Filter in +/-- If the `LSeries` of `f` and of `g` converge somewhere and agree on large real arguments, +then `f n = g n` whenever `n ≠ 0`. -/ +lemma LSeries.eq_of_LSeries_eventually_eq {f g : ℕ → ℂ} (hf : abscissaOfAbsConv f < ⊤) + (hg : abscissaOfAbsConv g < ⊤) (h : (fun x : ℝ ↦ LSeries f x) =ᶠ[atTop] fun x ↦ LSeries g x) + {n : ℕ} (hn : n ≠ 0) : + f n = g n := by + have hsub : (fun x : ℝ ↦ LSeries (f - g) x) =ᶠ[atTop] (0 : ℝ → ℂ) := + LSeries_sub_eventuallyEq_zero_of_LSeries_eventually_eq hf hg h + have ha : abscissaOfAbsConv (f - g) ≠ ⊤ := + lt_top_iff_ne_top.mp <| (abscissaOfAbsConv_sub_le f g).trans_lt <| max_lt hf hg + simpa only [Pi.sub_apply, sub_eq_zero] + using (LSeries_eventually_eq_zero_iff'.mp hsub).resolve_right ha n hn + +/-- If the `LSeries` of `f` and of `g` both converge somewhere, then they are equal if and only +if `f n = g n` whenever `n ≠ 0`. -/ +lemma LSeries_eq_iff_of_abscissaOfAbsConv_lt_top {f g : ℕ → ℂ} (hf : abscissaOfAbsConv f < ⊤) + (hg : abscissaOfAbsConv g < ⊤) : + LSeries f = LSeries g ↔ ∀ n ≠ 0, f n = g n := by + refine ⟨fun H n hn ↦ ?_, fun H ↦ funext (LSeries_congr · fun {n} ↦ H n)⟩ + refine eq_of_LSeries_eventually_eq hf hg ?_ hn + exact Filter.Eventually.of_forall fun x ↦ congr_fun H x + +/-- The map `f ↦ LSeries f` is injective on functions `f` such that `f 0 = 0` and the L-series +of `f` converges somewhere. -/ +lemma LSeries_injOn : Set.InjOn LSeries {f | f 0 = 0 ∧ abscissaOfAbsConv f < ⊤} := by + intro f hf g hg h + simp only [Set.mem_setOf] at hf hg + replace h := (LSeries_eq_iff_of_abscissaOfAbsConv_lt_top hf.2 hg.2).mp h + ext1 n + cases n with + | zero => exact hf.1.trans hg.1.symm + | succ n => exact h _ n.zero_ne_add_one.symm From 5d3e2fcb1bf9e75bb91636e1670095a09f49e7c4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 2 Jan 2025 12:06:36 +0000 Subject: [PATCH 078/189] feat(CategoryTheory/Sites): helper definitions for sieves (#19517) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For (pre)sieves constructed with `bind`, we introduce a structure `BindStruct` which contains the data and properties we may get when a morphism belongs to `Presieve.bind S R`. Similarly, we introduce definitions and lemmas `Sieve.ofArrows.i`, `Sieve.ofArrows.h`, `Sieve.ofArrows.fac` which provide a choice of factorization of a morphism which belong to a sieve generated by a family of arrows. This shall be necessary in #19444 as we sometimes need to choose such data in order to perform some other constructions. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib/CategoryTheory/Sites/Sieves.lean | 70 +++++++++++++++++++++--- 1 file changed, 63 insertions(+), 7 deletions(-) diff --git a/Mathlib/CategoryTheory/Sites/Sieves.lean b/Mathlib/CategoryTheory/Sites/Sieves.lean index 38b3981948c40..f2446cb5f188c 100644 --- a/Mathlib/CategoryTheory/Sites/Sieves.lean +++ b/Mathlib/CategoryTheory/Sites/Sieves.lean @@ -71,6 +71,34 @@ abbrev cocone (S : Presieve X) : Cocone S.diagram := def bind (S : Presieve X) (R : ∀ ⦃Y⦄ ⦃f : Y ⟶ X⦄, S f → Presieve Y) : Presieve X := fun Z h => ∃ (Y : C) (g : Z ⟶ Y) (f : Y ⟶ X) (H : S f), R H g ∧ g ≫ f = h +/-- Structure which contains the data and properties for a morphism `h` satisfying +`Presieve.bind S R h`. -/ +structure BindStruct (S : Presieve X) (R : ∀ ⦃Y⦄ ⦃f : Y ⟶ X⦄, S f → Presieve Y) + {Z : C} (h : Z ⟶ X) where + /-- the intermediate object -/ + Y : C + /-- a morphism in the family of presieves `R` -/ + g : Z ⟶ Y + /-- a morphism in the presieve `S` -/ + f : Y ⟶ X + hf : S f + hg : R hf g + fac : g ≫ f = h + +attribute [reassoc (attr := simp)] BindStruct.fac + +/-- If a morphism `h` satisfies `Presieve.bind S R h`, this is a choice of a structure +in `BindStruct S R h`. -/ +noncomputable def bind.bindStruct {S : Presieve X} {R : ∀ ⦃Y⦄ ⦃f : Y ⟶ X⦄, S f → Presieve Y} + {Z : C} {h : Z ⟶ X} (H : bind S R h) : BindStruct S R h := + Nonempty.some (by + obtain ⟨Y, g, f, hf, hg, fac⟩ := H + exact ⟨{ hf := hf, hg := hg, fac := fac }⟩) + +lemma BindStruct.bind {S : Presieve X} {R : ∀ ⦃Y⦄ ⦃f : Y ⟶ X⦄, S f → Presieve Y} + {Z : C} {h : Z ⟶ X} (b : BindStruct S R h) : bind S R h := + ⟨b.Y, b.g, b.f, b.hf, b.hg, b.fac⟩ + @[simp] theorem bind_comp {S : Presieve X} {R : ∀ ⦃Y : C⦄ ⦃f : Y ⟶ X⦄, S f → Presieve Y} {g : Z ⟶ Y} (h₁ : S f) (h₂ : R h₁ g) : bind S R (g ≫ f) := @@ -368,6 +396,12 @@ def bind (S : Presieve X) (R : ∀ ⦃Y⦄ ⦃f : Y ⟶ X⦄, S f → Sieve Y) : rintro Y Z f ⟨W, f, h, hh, hf, rfl⟩ g exact ⟨_, g ≫ f, _, hh, by simp [hf]⟩ +/-- Structure which contains the data and properties for a morphism `h` satisfying +`Sieve.bind S R h`. -/ +abbrev BindStruct (S : Presieve X) (R : ∀ ⦃Y⦄ ⦃f : Y ⟶ X⦄, S f → Sieve Y) + {Z : C} (h : Z ⟶ X) := + Presieve.BindStruct S (fun _ _ hf ↦ R hf) h + open Order Lattice theorem generate_le_iff (R : Presieve X) (S : Sieve X) : generate R ≤ S ↔ R ≤ S := @@ -417,17 +451,17 @@ lemma comp_mem_iff (i : X ⟶ Y) (f : Y ⟶ Z) [IsIso i] (S : Sieve Z) : convert S.downward_closed H (inv i) simp +section + +variable {I : Type*} {X : C} (Y : I → C) (f : ∀ i, Y i ⟶ X) + /-- The sieve of `X` generated by family of morphisms `Y i ⟶ X`. -/ -abbrev ofArrows {I : Type*} {X : C} (Y : I → C) (f : ∀ i, Y i ⟶ X) : - Sieve X := - generate (Presieve.ofArrows Y f) +abbrev ofArrows : Sieve X := generate (Presieve.ofArrows Y f) -lemma ofArrows_mk {I : Type*} {X : C} (Y : I → C) (f : ∀ i, Y i ⟶ X) (i : I) : - ofArrows Y f (f i) := +lemma ofArrows_mk (i : I) : ofArrows Y f (f i) := ⟨_, 𝟙 _, _, ⟨i⟩, by simp⟩ -lemma mem_ofArrows_iff {I : Type*} {X : C} (Y : I → C) (f : ∀ i, Y i ⟶ X) - {W : C} (g : W ⟶ X) : +lemma mem_ofArrows_iff {W : C} (g : W ⟶ X) : ofArrows Y f g ↔ ∃ (i : I) (a : W ⟶ Y i), g = a ≫ f i := by constructor · rintro ⟨T, a, b, ⟨i⟩, rfl⟩ @@ -435,6 +469,28 @@ lemma mem_ofArrows_iff {I : Type*} {X : C} (Y : I → C) (f : ∀ i, Y i ⟶ X) · rintro ⟨i, a, rfl⟩ apply downward_closed _ (ofArrows_mk Y f i) +variable {Y f} {W : C} {g : W ⟶ X} (hg : ofArrows Y f g) + +include hg in +lemma ofArrows.exists : ∃ (i : I) (h : W ⟶ Y i), g = h ≫ f i := by + obtain ⟨_, h, _, H, rfl⟩ := hg + cases' H with i + exact ⟨i, h, rfl⟩ + +/-- When `hg : Sieve.ofArrows Y f g`, this is a choice of `i` such that `g` +factors through `f i`. -/ +noncomputable def ofArrows.i : I := (ofArrows.exists hg).choose + +/-- When `hg : Sieve.ofArrows Y f g`, this is a morphism `h : W ⟶ Y (i hg)` such +that `h ≫ f (i hg) = g`. -/ +noncomputable def ofArrows.h : W ⟶ Y (i hg) := (ofArrows.exists hg).choose_spec.choose + +@[reassoc (attr := simp)] +lemma ofArrows.fac : h hg ≫ f (i hg) = g := + (ofArrows.exists hg).choose_spec.choose_spec.symm + +end + /-- The sieve generated by two morphisms. -/ abbrev ofTwoArrows {U V X : C} (i : U ⟶ X) (j : V ⟶ X) : Sieve X := Sieve.ofArrows (Y := pairFunction U V) (fun k ↦ WalkingPair.casesOn k i j) From 04bf741e0f2ed825214160021171d46bffefa2fd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= <6992158+datokrat@users.noreply.github.com> Date: Thu, 2 Jan 2025 12:06:38 +0000 Subject: [PATCH 079/189] feat(CategoryTheory/Sites): categories of sheaves are Grothendieck abelian (#19986) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If `J` is a Grothendieck topology on a small category `C : Type v`, and `A : Type u₁` (with `Category.{v} A`) is a Grothendieck abelian category, then `Sheaf J A` is a Grothendieck abelian category. Co-authored-by: Joël Riou Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> Co-authored-by: Paul Reichert --- .../Abelian/GrothendieckAxioms/Basic.lean | 43 +++++++++++- .../Abelian/GrothendieckAxioms/Sheaf.lean | 65 +++++++++++++++---- 2 files changed, 95 insertions(+), 13 deletions(-) diff --git a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Basic.lean b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Basic.lean index b96f54ec72e3b..5aaabd4f0aca1 100644 --- a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Basic.lean +++ b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Basic.lean @@ -53,7 +53,7 @@ open Limits attribute [instance] comp_preservesFiniteLimits comp_preservesFiniteColimits -universe w w' w₂ w₂' v v' u u' +universe w w' w₂ w₂' v v' v'' u u' u'' variable (C : Type u) [Category.{v} C] @@ -168,6 +168,47 @@ lemma HasExactLimitsOfShape.of_codomain_equivalence (J : Type*) [Category J] {D apply e.symm.congrRight.fullyFaithfulFunctor.preimageIso exact isoWhiskerLeft (_ ⋙ lim) e.unitIso.symm ≪≫ (preservesLimitNatIso e.inverse).symm +namespace Adjunction + +variable {C} {D : Type u''} [Category.{v''} D] {F : C ⥤ D} {G : D ⥤ C} + +/-- Let `adj : F ⊣ G` be an adjunction, with `G : D ⥤ C` reflective. +Assume that `D` has finite limits and `F` commutes to them. +If `C` has exact colimits of shape `J`, then `D` also has exact colimits of shape `J`. -/ +lemma hasExactColimitsOfShape (adj : F ⊣ G) [G.Full] [G.Faithful] + (J : Type u') [Category.{v'} J] [HasColimitsOfShape J C] [HasColimitsOfShape J D] + [HasExactColimitsOfShape J C] [HasFiniteLimits D] [PreservesFiniteLimits F] : + HasExactColimitsOfShape J D where + preservesFiniteLimits := ⟨fun K _ _ ↦ ⟨fun {H} ↦ by + have : PreservesLimitsOfSize.{0, 0} G := adj.rightAdjoint_preservesLimits + have : PreservesColimitsOfSize.{v', u'} F := adj.leftAdjoint_preservesColimits + let e : (whiskeringRight J D C).obj G ⋙ colim ⋙ F ≅ colim := + isoWhiskerLeft _ (preservesColimitNatIso F) ≪≫ (Functor.associator _ _ _).symm ≪≫ + isoWhiskerRight (whiskeringRightObjCompIso G F) _ ≪≫ + isoWhiskerRight ((whiskeringRight J D D).mapIso (asIso adj.counit)) _ ≪≫ + isoWhiskerRight wiskeringRightObjIdIso _ ≪≫ colim.leftUnitor + exact preservesLimit_of_natIso _ e⟩⟩ + +/-- Let `adj : F ⊣ G` be an adjunction, with `F : C ⥤ D` coreflective. +Assume that `C` has finite colimits and `G` commutes to them. +If `D` has exact limits of shape `J`, then `C` also has exact limits of shape `J`. -/ +lemma hasExactLimitsOfShape (adj : F ⊣ G) [F.Full] [F.Faithful] + (J : Type u') [Category.{v'} J] [HasLimitsOfShape J C] [HasLimitsOfShape J D] + [HasExactLimitsOfShape J D] [HasFiniteColimits C] [PreservesFiniteColimits G] : + HasExactLimitsOfShape J C where + preservesFiniteColimits:= ⟨fun K _ _ ↦ ⟨fun {H} ↦ by + have : PreservesLimitsOfSize.{v', u'} G := adj.rightAdjoint_preservesLimits + have : PreservesColimitsOfSize.{0, 0} F := adj.leftAdjoint_preservesColimits + let e : (whiskeringRight J _ _).obj F ⋙ lim ⋙ G ≅ lim := + isoWhiskerLeft _ (preservesLimitNatIso G) ≪≫ + (Functor.associator _ _ _).symm ≪≫ + isoWhiskerRight (whiskeringRightObjCompIso F G) _ ≪≫ + isoWhiskerRight ((whiskeringRight J C C).mapIso (asIso adj.unit).symm) _ ≪≫ + isoWhiskerRight wiskeringRightObjIdIso _ ≪≫ lim.leftUnitor + exact preservesColimit_of_natIso _ e⟩⟩ + +end Adjunction + /-- A category `C` which has coproducts is said to have `AB4` of size `w` provided that coproducts of size `w` are exact. diff --git a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Sheaf.lean b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Sheaf.lean index 7696bed4e68af..2cffda5728fe3 100644 --- a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Sheaf.lean +++ b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Sheaf.lean @@ -1,32 +1,73 @@ /- Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Dagur Asgeirsson +Authors: Dagur Asgeirsson, Joël Riou -/ import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms.FunctorCategory -import Mathlib.CategoryTheory.Sites.Limits +import Mathlib.CategoryTheory.Abelian.GrothendieckCategory +import Mathlib.CategoryTheory.Generator.Sheaf +import Mathlib.CategoryTheory.Sites.Abelian + /-! # AB axioms in sheaf categories -This file proves that, when the relevant limits and colimits and sheafification exist, exactness of -limits and colimits carries over from `A` to categories of `A`-valued sheaves. +If `J` is a Grothendieck topology on a small category `C : Type v`, +and `A : Type u₁` (with `Category.{v} A`) is a Grothendieck abelian category, +then `Sheaf J A` is a Grothendieck abelian category. + -/ +universe v v₁ v₂ u u₁ u₂ + namespace CategoryTheory open Limits -variable {A C J : Type*} [Category A] [Category C] [Category J] +namespace Sheaf + +variable {C : Type u} {A : Type u₁} {K : Type u₂} + [Category.{v} C] [Category.{v₁} A] [Category.{v₂} K] + (J : GrothendieckTopology C) + +section + +/- The two instances in this section apply in very rare situations, as they assume +that the forgetful functor from sheaves to presheaves commutes with certain colimits. +This does apply for sheaves for the extensive topology --- condensed modules over a +ring are examples of such sheaves. -/ + +variable [HasWeakSheafify J A] + +instance [HasFiniteLimits A] [HasColimitsOfShape K A] [HasExactColimitsOfShape K A] + [PreservesColimitsOfShape K (sheafToPresheaf J A)] : HasExactColimitsOfShape K (Sheaf J A) := + HasExactColimitsOfShape.domain_of_functor K (sheafToPresheaf J A) + +instance [HasFiniteColimits A] [HasLimitsOfShape K A] [HasExactLimitsOfShape K A] + [PreservesFiniteColimits (sheafToPresheaf J A)] : HasExactLimitsOfShape K (Sheaf J A) := + HasExactLimitsOfShape.domain_of_functor K (sheafToPresheaf J A) + +end + +instance hasFilteredColimitsOfSize + [HasSheafify J A] [HasFilteredColimitsOfSize.{v₂, u₂} A] : + HasFilteredColimitsOfSize.{v₂, u₂} (Sheaf J A) where + HasColimitsOfShape K := by infer_instance + +instance hasExactColimitsOfShape [HasFiniteLimits A] [HasSheafify J A] + [HasColimitsOfShape K A] [HasExactColimitsOfShape K A] : + HasExactColimitsOfShape K (Sheaf J A) := + (sheafificationAdjunction J A).hasExactColimitsOfShape K -variable (K : GrothendieckTopology C) [HasWeakSheafify K A] +instance ab5ofSize [HasFiniteLimits A] [HasSheafify J A] + [HasFilteredColimitsOfSize.{v₂, u₂} A] [AB5OfSize.{v₂, u₂} A] : + AB5OfSize.{v₂, u₂} (Sheaf J A) where + ofShape K _ _ := by infer_instance -instance [HasFiniteLimits A] [HasColimitsOfShape J A] [HasExactColimitsOfShape J A] - [PreservesColimitsOfShape J (sheafToPresheaf K A)] : HasExactColimitsOfShape J (Sheaf K A) := - HasExactColimitsOfShape.domain_of_functor J (sheafToPresheaf K A) +instance {C : Type v} [SmallCategory.{v} C] (J : GrothendieckTopology C) + (A : Type u₁) [Category.{v} A] [Abelian A] [IsGrothendieckAbelian.{v} A] + [HasSheafify J A] : IsGrothendieckAbelian.{v} (Sheaf J A) where -instance [HasFiniteColimits A] [HasLimitsOfShape J A] [HasExactLimitsOfShape J A] - [PreservesFiniteColimits (sheafToPresheaf K A)] : HasExactLimitsOfShape J (Sheaf K A) := - HasExactLimitsOfShape.domain_of_functor J (sheafToPresheaf K A) +end Sheaf end CategoryTheory From 4e190ae6c57bd874699b6896e3d483a1ed6228be Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 2 Jan 2025 12:06:39 +0000 Subject: [PATCH 080/189] feat(CategoryTheory): whiskering and curryfication of functors in three variables (#20197) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 1 + .../CategoryTheory/Functor/CurryingThree.lean | 100 ++++++++++++++++ .../CategoryTheory/Functor/Trifunctor.lean | 113 +++++++++++++++--- Mathlib/CategoryTheory/NatIso.lean | 16 ++- Mathlib/CategoryTheory/Whiskering.lean | 65 +++++++++- 5 files changed, 277 insertions(+), 18 deletions(-) create mode 100644 Mathlib/CategoryTheory/Functor/CurryingThree.lean diff --git a/Mathlib.lean b/Mathlib.lean index c07b91155b702..d20173600b13b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1767,6 +1767,7 @@ import Mathlib.CategoryTheory.Functor.Basic import Mathlib.CategoryTheory.Functor.Category import Mathlib.CategoryTheory.Functor.Const import Mathlib.CategoryTheory.Functor.Currying +import Mathlib.CategoryTheory.Functor.CurryingThree import Mathlib.CategoryTheory.Functor.Derived.RightDerived import Mathlib.CategoryTheory.Functor.EpiMono import Mathlib.CategoryTheory.Functor.Flat diff --git a/Mathlib/CategoryTheory/Functor/CurryingThree.lean b/Mathlib/CategoryTheory/Functor/CurryingThree.lean new file mode 100644 index 0000000000000..15628787f3423 --- /dev/null +++ b/Mathlib/CategoryTheory/Functor/CurryingThree.lean @@ -0,0 +1,100 @@ +/- +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.Functor.Currying +import Mathlib.CategoryTheory.Functor.Trifunctor +import Mathlib.CategoryTheory.Products.Associator + +/-! +# Currying of functors in three variables + +We study the equivalence of categories +`currying₃ : (C₁ ⥤ C₂ ⥤ C₃ ⥤ E) ≌ C₁ × C₂ × C₃ ⥤ E`. + +-/ + +namespace CategoryTheory + +variable {C₁ C₂ C₁₂ C₃ C₂₃ D₁ D₂ D₃ E : Type*} + [Category C₁] [Category C₂] [Category C₃] [Category C₁₂] [Category C₂₃] + [Category D₁] [Category D₂] [Category D₃] [Category E] + +/-- The equivalence of categories `(C₁ ⥤ C₂ ⥤ C₃ ⥤ E) ≌ C₁ × C₂ × C₃ ⥤ E` +given by the curryfication of functors in three variables. -/ +def currying₃ : (C₁ ⥤ C₂ ⥤ C₃ ⥤ E) ≌ C₁ × C₂ × C₃ ⥤ E := + currying.trans (currying.trans (prod.associativity C₁ C₂ C₃).congrLeft) + +/-- Uncurrying a functor in three variables. -/ +abbrev uncurry₃ : (C₁ ⥤ C₂ ⥤ C₃ ⥤ E) ⥤ C₁ × C₂ × C₃ ⥤ E := currying₃.functor + +/-- Currying a functor in three variables. -/ +abbrev curry₃ : (C₁ × C₂ × C₃ ⥤ E) ⥤ C₁ ⥤ C₂ ⥤ C₃ ⥤ E := currying₃.inverse + +/-- Uncurrying functors in three variables gives a fully faithful functor. -/ +def fullyFaithfulUncurry₃ : + (uncurry₃ : (C₁ ⥤ C₂ ⥤ C₃ ⥤ E) ⥤ (C₁ × C₂ × C₃ ⥤ E)).FullyFaithful := + currying₃.fullyFaithfulFunctor + +@[simp] +lemma curry₃_obj_map_app_app (F : C₁ × C₂ × C₃ ⥤ E) + {X₁ Y₁ : C₁} (f : X₁ ⟶ Y₁) (X₂ : C₂) (X₃ : C₃) : + (((curry₃.obj F).map f).app X₂).app X₃ = F.map ⟨f, 𝟙 X₂, 𝟙 X₃⟩ := rfl + +@[simp] +lemma curry₃_obj_obj_map_app (F : C₁ × C₂ × C₃ ⥤ E) + (X₁ : C₁) {X₂ Y₂ : C₂} (f : X₂ ⟶ Y₂) (X₃ : C₃) : + (((curry₃.obj F).obj X₁).map f).app X₃ = F.map ⟨𝟙 X₁, f, 𝟙 X₃⟩ := rfl + +@[simp] +lemma curry₃_obj_obj_obj_map (F : C₁ × C₂ × C₃ ⥤ E) + (X₁ : C₁) (X₂ : C₂) {X₃ Y₃ : C₃} (f : X₃ ⟶ Y₃) : + (((curry₃.obj F).obj X₁).obj X₂).map f = F.map ⟨𝟙 X₁, 𝟙 X₂, f⟩ := rfl + +@[simp] +lemma curry₃_map_app_app_app {F G : C₁ × C₂ × C₃ ⥤ E} (f : F ⟶ G) + (X₁ : C₁) (X₂ : C₂) (X₃ : C₃) : + (((curry₃.map f).app X₁).app X₂).app X₃ = f.app ⟨X₁, X₂, X₃⟩ := rfl + +@[simp] +lemma currying₃_unitIso_hom_app_app_app_app (F : C₁ ⥤ C₂ ⥤ C₃ ⥤ E) + (X₁ : C₁) (X₂ : C₂) (X₃ : C₃) : + (((currying₃.unitIso.hom.app F).app X₁).app X₂).app X₃ = 𝟙 _ := by + simp [currying₃, Equivalence.unit] + +@[simp] +lemma currying₃_unitIso_inv_app_app_app_app (F : C₁ ⥤ C₂ ⥤ C₃ ⥤ E) + (X₁ : C₁) (X₂ : C₂) (X₃ : C₃) : + (((currying₃.unitIso.inv.app F).app X₁).app X₂).app X₃ = 𝟙 _ := by + simp [currying₃, Equivalence.unitInv] + +/-- Given functors `F₁ : C₁ ⥤ D₁`, `F₂ : C₂ ⥤ D₂`, `F₃ : C₃ ⥤ D₃` +and `G : D₁ × D₂ × D₃ ⥤ E`, this is the isomorphism between +`curry₃.obj (F₁.prod (F₂.prod F₃) ⋙ G) : C₁ ⥤ C₂ ⥤ C₃ ⥤ E` +and `F₁ ⋙ curry₃.obj G ⋙ ((whiskeringLeft₂ E).obj F₂).obj F₃`. -/ +@[simps!] +def curry₃ObjProdComp (F₁ : C₁ ⥤ D₁) (F₂ : C₂ ⥤ D₂) (F₃ : C₃ ⥤ D₃) (G : D₁ × D₂ × D₃ ⥤ E) : + curry₃.obj (F₁.prod (F₂.prod F₃) ⋙ G) ≅ + F₁ ⋙ curry₃.obj G ⋙ ((whiskeringLeft₂ E).obj F₂).obj F₃ := + NatIso.ofComponents + (fun X₁ ↦ NatIso.ofComponents + (fun X₂ ↦ NatIso.ofComponents (fun X₃ ↦ Iso.refl _))) + +/-- `bifunctorComp₁₂` can be described in terms of the curryfication of functors. -/ +@[simps!] +def bifunctorComp₁₂Iso (F₁₂ : C₁ ⥤ C₂ ⥤ C₁₂) (G : C₁₂ ⥤ C₃ ⥤ E) : + bifunctorComp₁₂ F₁₂ G ≅ curry.obj (uncurry.obj F₁₂ ⋙ G) := + NatIso.ofComponents (fun _ => NatIso.ofComponents (fun _ => Iso.refl _)) + +/-- `bifunctorComp₂₃` can be described in terms of the curryfication of functors. -/ +@[simps!] +def bifunctorComp₂₃Iso (F : C₁ ⥤ C₂₃ ⥤ E) (G₂₃ : C₂ ⥤ C₃ ⥤ C₂₃) : + bifunctorComp₂₃ F G₂₃ ≅ + curry.obj (curry.obj (prod.associator _ _ _ ⋙ + uncurry.obj (uncurry.obj G₂₃ ⋙ F.flip).flip)) := + NatIso.ofComponents (fun _ ↦ NatIso.ofComponents (fun _ ↦ + NatIso.ofComponents (fun _ ↦ Iso.refl _))) + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Functor/Trifunctor.lean b/Mathlib/CategoryTheory/Functor/Trifunctor.lean index f1293cf7bb221..0548f4dadccdb 100644 --- a/Mathlib/CategoryTheory/Functor/Trifunctor.lean +++ b/Mathlib/CategoryTheory/Functor/Trifunctor.lean @@ -22,13 +22,12 @@ namespace CategoryTheory variable {C₁ C₂ C₃ C₄ C₁₂ C₂₃ : Type*} [Category C₁] [Category C₂] [Category C₃] [Category C₄] [Category C₁₂] [Category C₂₃] -section - -variable (F₁₂ : C₁ ⥤ C₂ ⥤ C₁₂) (G : C₁₂ ⥤ C₃ ⥤ C₄) +section bifunctorComp₁₂Functor /-- Auxiliary definition for `bifunctorComp₁₂`. -/ @[simps] -def bifunctorComp₁₂Obj (X₁ : C₁) : C₂ ⥤ C₃ ⥤ C₄ where +def bifunctorComp₁₂Obj (F₁₂ : C₁ ⥤ C₂ ⥤ C₁₂) (G : C₁₂ ⥤ C₃ ⥤ C₄) (X₁ : C₁) : + C₂ ⥤ C₃ ⥤ C₄ where obj X₂ := { obj := fun X₃ => (G.obj ((F₁₂.obj X₁).obj X₂)).obj X₃ map := fun {_ _} φ => (G.obj ((F₁₂.obj X₁).obj X₂)).map φ } @@ -38,7 +37,8 @@ def bifunctorComp₁₂Obj (X₁ : C₁) : C₂ ⥤ C₃ ⥤ C₄ where /-- Given two bifunctors `F₁₂ : C₁ ⥤ C₂ ⥤ C₁₂` and `G : C₁₂ ⥤ C₃ ⥤ C₄`, this is the trifunctor `C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄` obtained by composition. -/ @[simps] -def bifunctorComp₁₂ : C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄ where +def bifunctorComp₁₂ (F₁₂ : C₁ ⥤ C₂ ⥤ C₁₂) (G : C₁₂ ⥤ C₃ ⥤ C₄) : + C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄ where obj X₁ := bifunctorComp₁₂Obj F₁₂ G X₁ map {X₁ Y₁} φ := { app := fun X₂ => @@ -48,33 +48,116 @@ def bifunctorComp₁₂ : C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄ where dsimp simp only [← NatTrans.comp_app, ← G.map_comp, NatTrans.naturality] } -end +/-- Auxiliary definition for `bifunctorComp₁₂Functor`. -/ +@[simps] +def bifunctorComp₁₂FunctorObj (F₁₂ : C₁ ⥤ C₂ ⥤ C₁₂) : + (C₁₂ ⥤ C₃ ⥤ C₄) ⥤ C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄ where + obj G := bifunctorComp₁₂ F₁₂ G + map {G G'} φ := + { app X₁ := + { app X₂ := + { app X₃ := (φ.app ((F₁₂.obj X₁).obj X₂)).app X₃ } + naturality := fun X₂ Y₂ f ↦ by + ext X₃ + dsimp + simp only [← NatTrans.comp_app, NatTrans.naturality] } + naturality X₁ Y₁ f := by + ext X₂ X₃ + dsimp + simp only [← NatTrans.comp_app, NatTrans.naturality] } + +/-- Auxiliary definition for `bifunctorComp₁₂Functor`. -/ +@[simps] +def bifunctorComp₁₂FunctorMap {F₁₂ F₁₂' : C₁ ⥤ C₂ ⥤ C₁₂} (φ : F₁₂ ⟶ F₁₂') : + bifunctorComp₁₂FunctorObj (C₃ := C₃) (C₄ := C₄) F₁₂ ⟶ bifunctorComp₁₂FunctorObj F₁₂' where + app G := + { app X₁ := + { app X₂ := { app X₃ := (G.map ((φ.app X₁).app X₂)).app X₃ } + naturality := fun X₂ Y₂ f ↦ by + ext X₃ + dsimp + simp only [← NatTrans.comp_app, NatTrans.naturality, ← G.map_comp] } + naturality X₁ Y₁ f := by + ext X₂ X₃ + dsimp + simp only [← NatTrans.comp_app, NatTrans.naturality, ← G.map_comp] } + naturality G G' f := by + ext X₁ X₂ X₃ + dsimp + simp only [← NatTrans.comp_app, NatTrans.naturality] + +/-- The functor `(C₁ ⥤ C₂ ⥤ C₁₂) ⥤ (C₁₂ ⥤ C₃ ⥤ C₄) ⥤ C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄` which +sends `F₁₂ : C₁ ⥤ C₂ ⥤ C₁₂` and `G : C₁₂ ⥤ C₃ ⥤ C₄` to the functor +`bifunctorComp₁₂ F₁₂ G : C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄`. -/ +@[simps] +def bifunctorComp₁₂Functor : (C₁ ⥤ C₂ ⥤ C₁₂) ⥤ (C₁₂ ⥤ C₃ ⥤ C₄) ⥤ C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄ where + obj := bifunctorComp₁₂FunctorObj + map := bifunctorComp₁₂FunctorMap -section +end bifunctorComp₁₂Functor -variable (F : C₁ ⥤ C₂₃ ⥤ C₄) (G₂₃ : C₂ ⥤ C₃ ⥤ C₂₃) +section bifunctorComp₂₃Functor /-- Auxiliary definition for `bifunctorComp₂₃`. -/ @[simps] -def bifunctorComp₂₃Obj (X₁ : C₁) : C₂ ⥤ C₃ ⥤ C₄ where +def bifunctorComp₂₃Obj (F : C₁ ⥤ C₂₃ ⥤ C₄) (G₂₃ : C₂ ⥤ C₃ ⥤ C₂₃) (X₁ : C₁) : + C₂ ⥤ C₃ ⥤ C₄ where obj X₂ := - { obj := fun X₃ => (F.obj X₁).obj ((G₂₃.obj X₂).obj X₃) - map := fun {_ _} φ => (F.obj X₁).map ((G₂₃.obj X₂).map φ) } + { obj X₃ := (F.obj X₁).obj ((G₂₃.obj X₂).obj X₃) + map φ := (F.obj X₁).map ((G₂₃.obj X₂).map φ) } map {X₂ Y₂} φ := - { app := fun X₃ => (F.obj X₁).map ((G₂₃.map φ).app X₃) - naturality := fun {X₃ Y₃} φ => by + { app X₃ := (F.obj X₁).map ((G₂₃.map φ).app X₃) + naturality X₃ Y₃ φ := by dsimp simp only [← Functor.map_comp, NatTrans.naturality] } /-- Given two bifunctors `F : C₁ ⥤ C₂₃ ⥤ C₄` and `G₂₃ : C₂ ⥤ C₃ ⥤ C₄`, this is the trifunctor `C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄` obtained by composition. -/ @[simps] -def bifunctorComp₂₃ : C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄ where +def bifunctorComp₂₃ (F : C₁ ⥤ C₂₃ ⥤ C₄) (G₂₃ : C₂ ⥤ C₃ ⥤ C₂₃) : + C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄ where obj X₁ := bifunctorComp₂₃Obj F G₂₃ X₁ map {X₁ Y₁} φ := { app := fun X₂ => { app := fun X₃ => (F.map φ).app ((G₂₃.obj X₂).obj X₃) } } -end +/-- Auxiliary definition for `bifunctorComp₂₃Functor`. -/ +@[simps] +def bifunctorComp₂₃FunctorObj (F : C₁ ⥤ C₂₃ ⥤ C₄) : + (C₂ ⥤ C₃ ⥤ C₂₃) ⥤ C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄ where + obj G₂₃ := bifunctorComp₂₃ F G₂₃ + map {G₂₃ G₂₃'} φ := + { app X₁ := + { app X₂ := + { app X₃ := (F.obj X₁).map ((φ.app X₂).app X₃) + naturality X₃ Y₃ f := by + dsimp + simp only [← Functor.map_comp, NatTrans.naturality] } + naturality X₂ Y₂ f := by + ext X₃ + dsimp + simp only [← NatTrans.comp_app, ← Functor.map_comp, NatTrans.naturality] } } + +/-- Auxiliary definition for `bifunctorComp₂₃Functor`. -/ +@[simps] +def bifunctorComp₂₃FunctorMap {F F' : C₁ ⥤ C₂₃ ⥤ C₄} (φ : F ⟶ F') : + bifunctorComp₂₃FunctorObj F (C₂ := C₂) (C₃ := C₃) ⟶ bifunctorComp₂₃FunctorObj F' where + app G₂₃ := + { app X₁ := { app X₂ := { app X₃ := (φ.app X₁).app ((G₂₃.obj X₂).obj X₃) } } + naturality X₁ Y₁ f := by + ext X₂ X₃ + dsimp + simp only [← NatTrans.comp_app, NatTrans.naturality] } + +/-- The functor `(C₁ ⥤ C₂₃ ⥤ C₄) ⥤ (C₂ ⥤ C₃ ⥤ C₂₃) ⥤ C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄` which +sends `F : C₁ ⥤ C₂₃ ⥤ C₄` and `G₂₃ : C₂ ⥤ C₃ ⥤ C₂₃` to the +functor `bifunctorComp₂₃ F G₂₃ : C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄`. -/ +@[simps] +def bifunctorComp₂₃Functor : + (C₁ ⥤ C₂₃ ⥤ C₄) ⥤ (C₂ ⥤ C₃ ⥤ C₂₃) ⥤ C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄ where + obj := bifunctorComp₂₃FunctorObj + map := bifunctorComp₂₃FunctorMap + +end bifunctorComp₂₃Functor end CategoryTheory diff --git a/Mathlib/CategoryTheory/NatIso.lean b/Mathlib/CategoryTheory/NatIso.lean index fcf0a2f963969..e12d02fb99a67 100644 --- a/Mathlib/CategoryTheory/NatIso.lean +++ b/Mathlib/CategoryTheory/NatIso.lean @@ -41,7 +41,7 @@ namespace CategoryTheory open NatTrans variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] {E : Type u₃} - [Category.{v₃} E] + [Category.{v₃} E] {E' : Type u₄} [Category.{v₄} E'] namespace Iso @@ -75,6 +75,20 @@ lemma inv_hom_id_app_app {F G : C ⥤ D ⥤ E} (e : F ≅ G) (X₁ : C) (X₂ : (e.inv.app X₁).app X₂ ≫ (e.hom.app X₁).app X₂ = 𝟙 _ := by rw [← NatTrans.comp_app, Iso.inv_hom_id_app, NatTrans.id_app] +@[reassoc (attr := simp)] +lemma hom_inv_id_app_app_app {F G : C ⥤ D ⥤ E ⥤ E'} (e : F ≅ G) + (X₁ : C) (X₂ : D) (X₃ : E) : + ((e.hom.app X₁).app X₂).app X₃ ≫ ((e.inv.app X₁).app X₂).app X₃ = 𝟙 _ := by + rw [← NatTrans.comp_app, ← NatTrans.comp_app, Iso.hom_inv_id_app, + NatTrans.id_app, NatTrans.id_app] + +@[reassoc (attr := simp)] +lemma inv_hom_id_app_app_app {F G : C ⥤ D ⥤ E ⥤ E'} (e : F ≅ G) + (X₁ : C) (X₂ : D) (X₃ : E) : + ((e.inv.app X₁).app X₂).app X₃ ≫ ((e.hom.app X₁).app X₂).app X₃ = 𝟙 _ := by + rw [← NatTrans.comp_app, ← NatTrans.comp_app, Iso.inv_hom_id_app, + NatTrans.id_app, NatTrans.id_app] + end Iso namespace NatIso diff --git a/Mathlib/CategoryTheory/Whiskering.lean b/Mathlib/CategoryTheory/Whiskering.lean index 53feb19a7b2f9..6983922d441e3 100644 --- a/Mathlib/CategoryTheory/Whiskering.lean +++ b/Mathlib/CategoryTheory/Whiskering.lean @@ -302,8 +302,9 @@ theorem pentagon : end Functor -variable {C₁ C₂ D₁ D₂ : Type*} [Category C₁] [Category C₂] - [Category D₁] [Category D₂] (E : Type*) [Category E] +variable {C₁ C₂ C₃ D₁ D₂ D₃ : Type*} [Category C₁] [Category C₂] [Category C₃] + [Category D₁] [Category D₂] [Category D₃] (E : Type*) [Category E] + /-- The obvious functor `(C₁ ⥤ D₁) ⥤ (C₂ ⥤ D₂) ⥤ (D₁ ⥤ D₂ ⥤ E) ⥤ (C₁ ⥤ C₂ ⥤ E)`. -/ @[simps!] @@ -318,4 +319,64 @@ def whiskeringLeft₂ : map ψ := { app := fun F₂ ↦ whiskerLeft _ ((whiskeringLeft C₁ D₁ (C₂ ⥤ E)).map ψ) } +/-- Auxiliary definition for `whiskeringLeft₃`. -/ +@[simps!] +def whiskeringLeft₃ObjObjObj (F₁ : C₁ ⥤ D₁) (F₂ : C₂ ⥤ D₂) (F₃ : C₃ ⥤ D₃) : + (D₁ ⥤ D₂ ⥤ D₃ ⥤ E) ⥤ C₁ ⥤ C₂ ⥤ C₃ ⥤ E := + (whiskeringRight _ _ _).obj (((whiskeringLeft₂ E).obj F₂).obj F₃) ⋙ + (whiskeringLeft C₁ D₁ _).obj F₁ + +/-- Auxiliary definition for `whiskeringLeft₃`. -/ +@[simps] +def whiskeringLeft₃ObjObjMap (F₁ : C₁ ⥤ D₁) (F₂ : C₂ ⥤ D₂) {F₃ F₃' : C₃ ⥤ D₃} (τ₃ : F₃ ⟶ F₃') : + whiskeringLeft₃ObjObjObj E F₁ F₂ F₃ ⟶ + whiskeringLeft₃ObjObjObj E F₁ F₂ F₃' where + app F := whiskerLeft _ (whiskerLeft _ (((whiskeringLeft₂ E).obj F₂).map τ₃)) + +variable (C₃ D₃) in +/-- Auxiliary definition for `whiskeringLeft₃`. -/ +@[simps] +def whiskeringLeft₃ObjObj (F₁ : C₁ ⥤ D₁) (F₂ : C₂ ⥤ D₂) : + (C₃ ⥤ D₃) ⥤ (D₁ ⥤ D₂ ⥤ D₃ ⥤ E) ⥤ (C₁ ⥤ C₂ ⥤ C₃ ⥤ E) where + obj F₃ := whiskeringLeft₃ObjObjObj E F₁ F₂ F₃ + map τ₃ := whiskeringLeft₃ObjObjMap E F₁ F₂ τ₃ + +variable (C₃ D₃) in +/-- Auxiliary definition for `whiskeringLeft₃`. -/ +@[simps] +def whiskeringLeft₃ObjMap (F₁ : C₁ ⥤ D₁) {F₂ F₂' : C₂ ⥤ D₂} (τ₂ : F₂ ⟶ F₂') : + whiskeringLeft₃ObjObj C₃ D₃ E F₁ F₂ ⟶ whiskeringLeft₃ObjObj C₃ D₃ E F₁ F₂' where + app F₃ := whiskerRight ((whiskeringRight _ _ _).map (((whiskeringLeft₂ E).map τ₂).app F₃)) _ + +variable (C₂ C₃ D₂ D₃) in +/-- Auxiliary definition for `whiskeringLeft₃`. -/ +@[simps] +def whiskeringLeft₃Obj (F₁ : C₁ ⥤ D₁) : + (C₂ ⥤ D₂) ⥤ (C₃ ⥤ D₃) ⥤ (D₁ ⥤ D₂ ⥤ D₃ ⥤ E) ⥤ (C₁ ⥤ C₂ ⥤ C₃ ⥤ E) where + obj F₂ := whiskeringLeft₃ObjObj C₃ D₃ E F₁ F₂ + map τ₂ := whiskeringLeft₃ObjMap C₃ D₃ E F₁ τ₂ + +variable (C₂ C₃ D₂ D₃) in +/-- Auxiliary definition for `whiskeringLeft₃`. -/ +@[simps] +def whiskeringLeft₃Map {F₁ F₁' : C₁ ⥤ D₁} (τ₁ : F₁ ⟶ F₁') : + whiskeringLeft₃Obj C₂ C₃ D₂ D₃ E F₁ ⟶ whiskeringLeft₃Obj C₂ C₃ D₂ D₃ E F₁' where + app F₂ := { app F₃ := whiskerLeft _ ((whiskeringLeft _ _ _).map τ₁) } + +/-- The obvious functor `(C₁ ⥤ D₁) ⥤ (C₂ ⥤ D₂) ⥤ (D₁ ⥤ D₂ ⥤ E) ⥤ (C₁ ⥤ C₂ ⥤ E)`. -/ +@[simps!] +def whiskeringLeft₃ : + (C₁ ⥤ D₁) ⥤ (C₂ ⥤ D₂) ⥤ (C₃ ⥤ D₃) ⥤ (D₁ ⥤ D₂ ⥤ D₃ ⥤ E) ⥤ (C₁ ⥤ C₂ ⥤ C₃ ⥤ E) where + obj F₁ := whiskeringLeft₃Obj C₂ C₃ D₂ D₃ E F₁ + map τ₁ := whiskeringLeft₃Map C₂ C₃ D₂ D₃ E τ₁ + +variable {E} in +/-- The "postcomposition" with a functor `E ⥤ E'` gives a functor +`(E ⥤ E') ⥤ (C₁ ⥤ C₂ ⥤ C₃ ⥤ E) ⥤ C₁ ⥤ C₂ ⥤ C₃ ⥤ E'`. -/ +@[simps!] +def Functor.postcompose₃ {E' : Type*} [Category E'] : + (E ⥤ E') ⥤ (C₁ ⥤ C₂ ⥤ C₃ ⥤ E) ⥤ C₁ ⥤ C₂ ⥤ C₃ ⥤ E' := + whiskeringRight C₃ _ _ ⋙ whiskeringRight C₂ _ _ ⋙ whiskeringRight C₁ _ _ + end CategoryTheory + From cbf0869ec450a1ecb6aaed922b60319098482f06 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 2 Jan 2025 12:06:41 +0000 Subject: [PATCH 081/189] chore(EventuallyMeasurable): move to `MeasurableSpace/` (#20213) I'm gradually moving material that doesn't depend on measures to `MeasurableSpace/`. --- Mathlib.lean | 2 +- .../EventuallyMeasurable.lean | 0 Mathlib/MeasureTheory/Measure/NullMeasurable.lean | 2 +- Mathlib/Topology/Baire/BaireMeasurable.lean | 2 +- 4 files changed, 3 insertions(+), 3 deletions(-) rename Mathlib/MeasureTheory/{Constructions => MeasurableSpace}/EventuallyMeasurable.lean (100%) diff --git a/Mathlib.lean b/Mathlib.lean index d20173600b13b..c15d73c88f662 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3631,7 +3631,6 @@ import Mathlib.MeasureTheory.Constructions.BorelSpace.Metrizable import Mathlib.MeasureTheory.Constructions.BorelSpace.Order import Mathlib.MeasureTheory.Constructions.BorelSpace.Real import Mathlib.MeasureTheory.Constructions.Cylinders -import Mathlib.MeasureTheory.Constructions.EventuallyMeasurable import Mathlib.MeasureTheory.Constructions.HaarToSphere import Mathlib.MeasureTheory.Constructions.Pi import Mathlib.MeasureTheory.Constructions.Polish.Basic @@ -3748,6 +3747,7 @@ import Mathlib.MeasureTheory.MeasurableSpace.Card import Mathlib.MeasureTheory.MeasurableSpace.CountablyGenerated import Mathlib.MeasureTheory.MeasurableSpace.Defs import Mathlib.MeasureTheory.MeasurableSpace.Embedding +import Mathlib.MeasureTheory.MeasurableSpace.EventuallyMeasurable import Mathlib.MeasureTheory.MeasurableSpace.Instances import Mathlib.MeasureTheory.MeasurableSpace.Invariants import Mathlib.MeasureTheory.MeasurableSpace.NCard diff --git a/Mathlib/MeasureTheory/Constructions/EventuallyMeasurable.lean b/Mathlib/MeasureTheory/MeasurableSpace/EventuallyMeasurable.lean similarity index 100% rename from Mathlib/MeasureTheory/Constructions/EventuallyMeasurable.lean rename to Mathlib/MeasureTheory/MeasurableSpace/EventuallyMeasurable.lean diff --git a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean index adefe7bb1eecd..b2e786979552e 100644 --- a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov -/ -import Mathlib.MeasureTheory.Constructions.EventuallyMeasurable +import Mathlib.MeasureTheory.MeasurableSpace.EventuallyMeasurable import Mathlib.MeasureTheory.MeasurableSpace.Basic import Mathlib.MeasureTheory.Measure.AEDisjoint diff --git a/Mathlib/Topology/Baire/BaireMeasurable.lean b/Mathlib/Topology/Baire/BaireMeasurable.lean index 86cceb1e9825f..1288e892880ac 100644 --- a/Mathlib/Topology/Baire/BaireMeasurable.lean +++ b/Mathlib/Topology/Baire/BaireMeasurable.lean @@ -5,7 +5,7 @@ Authors: Felix Weilacher -/ import Mathlib.Topology.GDelta.UniformSpace import Mathlib.Topology.LocallyClosed -import Mathlib.MeasureTheory.Constructions.EventuallyMeasurable +import Mathlib.MeasureTheory.MeasurableSpace.EventuallyMeasurable import Mathlib.MeasureTheory.Constructions.BorelSpace.Basic /-! From 8cef41c850692b72d6c8d4011affcff3dd301c9c Mon Sep 17 00:00:00 2001 From: Yoh Tanimoto Date: Thu, 2 Jan 2025 12:06:42 +0000 Subject: [PATCH 082/189] feat(MeasureTheory/Integral/SetIntegral): add the restriction of an integral to the support (#20258) add `MeasureTheory.integral_tsupport ` showing that an integral is equal to its restriction to the support of the integrated function. motivation: this will be used in #12290 to reduce the integral to the support and enables an estimate by the measure of `tsupport` and the sup of the function. Co-authored-by: Yoh Tanimoto <57562556+yoh-tanimoto@users.noreply.github.com> --- .../MeasureTheory/Integral/SetIntegral.lean | 18 ++++++++++++++++++ 1 file changed, 18 insertions(+) diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 8633fe486564e..f97aa2469f9e8 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -1072,6 +1072,24 @@ theorem Continuous.integral_pos_of_hasCompactSupport_nonneg_nonzero [IsFiniteMea end OpenPos +section Support + +variable {M : Type*} [NormedAddCommGroup M] [NormedSpace ℝ M] {mX : MeasurableSpace X} + {ν : Measure X} {F : X → M} + +theorem MeasureTheory.setIntegral_support : ∫ x in support F, F x ∂ν = ∫ x, F x ∂ν := by + nth_rw 2 [← setIntegral_univ] + rw [setIntegral_eq_of_subset_of_forall_diff_eq_zero MeasurableSet.univ (subset_univ (support F))] + exact fun _ hx => nmem_support.mp <| not_mem_of_mem_diff hx + +theorem MeasureTheory.setIntegral_tsupport [TopologicalSpace X] : + ∫ x in tsupport F, F x ∂ν = ∫ x, F x ∂ν := by + nth_rw 2 [← setIntegral_univ] + rw [setIntegral_eq_of_subset_of_forall_diff_eq_zero MeasurableSet.univ (subset_univ (tsupport F))] + exact fun _ hx => image_eq_zero_of_nmem_tsupport <| not_mem_of_mem_diff hx + +end Support + /-! Fundamental theorem of calculus for set integrals -/ section FTC From f52e49b5dbf092b92990875ff640f32130d440ac Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 2 Jan 2025 12:06:43 +0000 Subject: [PATCH 083/189] feat(Topology/Algebra): compact hausdorff fields are finite (#20289) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib/Topology/Algebra/Field.lean | 8 +++++++ Mathlib/Topology/Algebra/Monoid.lean | 35 ++++++++++++++++++++++++++++ 2 files changed, 43 insertions(+) diff --git a/Mathlib/Topology/Algebra/Field.lean b/Mathlib/Topology/Algebra/Field.lean index 665d8fe1857c3..cce773d29a6f3 100644 --- a/Mathlib/Topology/Algebra/Field.lean +++ b/Mathlib/Topology/Algebra/Field.lean @@ -32,6 +32,14 @@ theorem Filter.tendsto_cocompact_mul_right₀ [ContinuousMul K] {a : K} (ha : a Filter.Tendsto (fun x : K => x * a) (Filter.cocompact K) (Filter.cocompact K) := Filter.tendsto_cocompact_mul_right (mul_inv_cancel₀ ha) +/-- Compact hausdorff topological fields are finite. -/ +instance (priority := 100) {K} [DivisionRing K] [TopologicalSpace K] + [TopologicalRing K] [CompactSpace K] [T2Space K] : Finite K := by + suffices DiscreteTopology K by + exact finite_of_compact_of_discrete + rw [discreteTopology_iff_isOpen_singleton_zero] + exact GroupWithZero.isOpen_singleton_zero + variable (K) /-- A topological division ring is a division ring with a topology where all operations are diff --git a/Mathlib/Topology/Algebra/Monoid.lean b/Mathlib/Topology/Algebra/Monoid.lean index 8a4269b811165..ea4a62c0da6b2 100644 --- a/Mathlib/Topology/Algebra/Monoid.lean +++ b/Mathlib/Topology/Algebra/Monoid.lean @@ -406,6 +406,25 @@ open Filter variable {α β : Type*} variable [TopologicalSpace M] [MulZeroClass M] [ContinuousMul M] +theorem exists_mem_nhds_zero_mul_subset + {K U : Set M} (hK : IsCompact K) (hU : U ∈ 𝓝 0) : ∃ V ∈ 𝓝 0, K * V ⊆ U := by + refine hK.induction_on ?_ ?_ ?_ ?_ + · exact ⟨univ, by simp⟩ + · rintro s t hst ⟨V, hV, hV'⟩ + exact ⟨V, hV, (mul_subset_mul_right hst).trans hV'⟩ + · rintro s t ⟨V, V_in, hV'⟩ ⟨W, W_in, hW'⟩ + use V ∩ W, inter_mem V_in W_in + rw [union_mul] + exact + union_subset ((mul_subset_mul_left V.inter_subset_left).trans hV') + ((mul_subset_mul_left V.inter_subset_right).trans hW') + · intro x hx + have := tendsto_mul (show U ∈ 𝓝 (x * 0) by simpa using hU) + rw [nhds_prod_eq, mem_map, mem_prod_iff] at this + rcases this with ⟨t, ht, s, hs, h⟩ + rw [← image_subset_iff, image_mul_prod] at h + exact ⟨t, mem_nhdsWithin_of_mem_nhds ht, s, hs, h⟩ + /-- Let `M` be a topological space with a continuous multiplication operation and a `0`. Let `l` be a filter on `M` which is disjoint from the cocompact filter. Then, the multiplication map `M × M → M` tends to zero on the filter product `𝓝 0 ×ˢ l`. -/ @@ -507,6 +526,22 @@ theorem tendsto_mul_cofinite_nhds_zero {f : α → M} {g : β → M} end MulZeroClass +section GroupWithZero + +lemma GroupWithZero.isOpen_singleton_zero [GroupWithZero M] [TopologicalSpace M] + [ContinuousMul M] [CompactSpace M] [T1Space M] : + IsOpen {(0 : M)} := by + obtain ⟨U, hU, h0U, h1U⟩ := t1Space_iff_exists_open.mp ‹_› zero_ne_one + obtain ⟨W, hW, hW'⟩ := exists_mem_nhds_zero_mul_subset isCompact_univ (hU.mem_nhds h0U) + by_cases H : ∃ x ≠ 0, x ∈ W + · obtain ⟨x, hx, hxW⟩ := H + cases h1U (hW' (by simpa [hx] using Set.mul_mem_mul (Set.mem_univ x⁻¹) hxW)) + · obtain rfl : W = {0} := subset_antisymm + (by simpa [not_imp_not] using H) (by simpa using mem_of_mem_nhds hW) + simpa [isOpen_iff_mem_nhds] + +end GroupWithZero + section MulOneClass variable [TopologicalSpace M] [MulOneClass M] [ContinuousMul M] From f65e1af0ea5fa2ed56ae584d959d64f9059a76aa Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 2 Jan 2025 12:06:44 +0000 Subject: [PATCH 084/189] feat: `Fin.cons` and `Fin.snoc` are continuous (#20352) These are adapted from the results for `insertNth`, which now appear immediately below the new results. I've also: * extracted `section`s with appropriate `variables`s for these `Fin` results. * renamed from `fin_insertNth` to `finInsertNth`, to match the naming rules * reordered the parameters to put all the continuity assumptions at the end * not tagged these with `@[continuity]` as it doesn't seem to help solve `Continuous fun x : X => ![x, x]`. We can look at this in a follow-up. Co-authored-by: Johan Commelin --- Mathlib/Analysis/BoxIntegral/Box/Basic.lean | 2 +- Mathlib/Topology/Constructions.lean | 84 ++++++++++++++++++--- Mathlib/Topology/ContinuousOn.lean | 64 +++++++++++++--- 3 files changed, 128 insertions(+), 22 deletions(-) diff --git a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean index 1700dfa697aab..29cee2b8e130b 100644 --- a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean @@ -383,7 +383,7 @@ theorem continuousOn_face_Icc {X} [TopologicalSpace X] {n} {f : (Fin (n + 1) → {I : Box (Fin (n + 1))} (h : ContinuousOn f (Box.Icc I)) {i : Fin (n + 1)} {x : ℝ} (hx : x ∈ Icc (I.lower i) (I.upper i)) : ContinuousOn (f ∘ i.insertNth x) (Box.Icc (I.face i)) := - h.comp (continuousOn_const.fin_insertNth i continuousOn_id) (I.mapsTo_insertNth_face_Icc hx) + h.comp (continuousOn_const.finInsertNth i continuousOn_id) (I.mapsTo_insertNth_face_Icc hx) /-! ### Covering of the interior of a box by a monotone sequence of smaller boxes diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 97a7797155a00..ee4d5a39eb471 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot -/ import Mathlib.Data.Finset.Piecewise +import Mathlib.Data.Fin.VecNotation import Mathlib.Order.Filter.Curry import Mathlib.Topology.Maps.Basic import Mathlib.Topology.NhdsSet @@ -1416,22 +1417,83 @@ theorem continuous_mulSingle [∀ i, One (π i)] [DecidableEq ι] (i : ι) : Continuous fun x => (Pi.mulSingle i x : ∀ i, π i) := continuous_const.update _ continuous_id -theorem Filter.Tendsto.fin_insertNth {n} {π : Fin (n + 1) → Type*} [∀ i, TopologicalSpace (π i)] - (i : Fin (n + 1)) {f : Y → π i} {l : Filter Y} {x : π i} (hf : Tendsto f l (𝓝 x)) - {g : Y → ∀ j : Fin n, π (i.succAbove j)} {y : ∀ j, π (i.succAbove j)} (hg : Tendsto g l (𝓝 y)) : +section Fin +variable {n : ℕ} {π : Fin (n + 1) → Type*} [∀ i, TopologicalSpace (π i)] + +theorem Filter.Tendsto.finCons + {f : Y → π 0} {g : Y → ∀ j : Fin n, π j.succ} {l : Filter Y} {x : π 0} {y : ∀ j, π (Fin.succ j)} + (hf : Tendsto f l (𝓝 x)) (hg : Tendsto g l (𝓝 y)) : + Tendsto (fun a => Fin.cons (f a) (g a)) l (𝓝 <| Fin.cons x y) := + tendsto_pi_nhds.2 fun j => Fin.cases (by simpa) (by simpa using tendsto_pi_nhds.1 hg) j + +theorem ContinuousAt.finCons {f : X → π 0} {g : X → ∀ j : Fin n, π (Fin.succ j)} {x : X} + (hf : ContinuousAt f x) (hg : ContinuousAt g x) : + ContinuousAt (fun a => Fin.cons (f a) (g a)) x := + hf.tendsto.finCons hg + +theorem Continuous.finCons {f : X → π 0} {g : X → ∀ j : Fin n, π (Fin.succ j)} + (hf : Continuous f) (hg : Continuous g) : Continuous fun a => Fin.cons (f a) (g a) := + continuous_iff_continuousAt.2 fun _ => hf.continuousAt.finCons hg.continuousAt + +theorem Filter.Tendsto.matrixVecCons + {f : Y → Z} {g : Y → Fin n → Z} {l : Filter Y} {x : Z} {y : Fin n → Z} + (hf : Tendsto f l (𝓝 x)) (hg : Tendsto g l (𝓝 y)) : + Tendsto (fun a => Matrix.vecCons (f a) (g a)) l (𝓝 <| Matrix.vecCons x y) := + hf.finCons hg + +theorem ContinuousAt.matrixVecCons + {f : X → Z} {g : X → Fin n → Z} {x : X} (hf : ContinuousAt f x) (hg : ContinuousAt g x) : + ContinuousAt (fun a => Matrix.vecCons (f a) (g a)) x := + hf.finCons hg + +theorem Continuous.matrixVecCons + {f : X → Z} {g : X → Fin n → Z} (hf : Continuous f) (hg : Continuous g) : + Continuous fun a => Matrix.vecCons (f a) (g a) := + hf.finCons hg + +theorem Filter.Tendsto.finSnoc + {f : Y → ∀ j : Fin n, π j.castSucc} {g : Y → π (Fin.last _)} + {l : Filter Y} {x : ∀ j, π (Fin.castSucc j)} {y : π (Fin.last _)} + (hf : Tendsto f l (𝓝 x)) (hg : Tendsto g l (𝓝 y)) : + Tendsto (fun a => Fin.snoc (f a) (g a)) l (𝓝 <| Fin.snoc x y) := + tendsto_pi_nhds.2 fun j => Fin.lastCases (by simpa) (by simpa using tendsto_pi_nhds.1 hf) j + +theorem ContinuousAt.finSnoc {f : X → ∀ j : Fin n, π j.castSucc} {g : X → π (Fin.last _)} {x : X} + (hf : ContinuousAt f x) (hg : ContinuousAt g x) : + ContinuousAt (fun a => Fin.snoc (f a) (g a)) x := + hf.tendsto.finSnoc hg + +theorem Continuous.finSnoc {f : X → ∀ j : Fin n, π j.castSucc} {g : X → π (Fin.last _)} + (hf : Continuous f) (hg : Continuous g) : Continuous fun a => Fin.snoc (f a) (g a) := + continuous_iff_continuousAt.2 fun _ => hf.continuousAt.finSnoc hg.continuousAt + +theorem Filter.Tendsto.finInsertNth + (i : Fin (n + 1)) {f : Y → π i} {g : Y → ∀ j : Fin n, π (i.succAbove j)} {l : Filter Y} + {x : π i} {y : ∀ j, π (i.succAbove j)} (hf : Tendsto f l (𝓝 x)) (hg : Tendsto g l (𝓝 y)) : Tendsto (fun a => i.insertNth (f a) (g a)) l (𝓝 <| i.insertNth x y) := tendsto_pi_nhds.2 fun j => Fin.succAboveCases i (by simpa) (by simpa using tendsto_pi_nhds.1 hg) j -theorem ContinuousAt.fin_insertNth {n} {π : Fin (n + 1) → Type*} [∀ i, TopologicalSpace (π i)] - (i : Fin (n + 1)) {f : X → π i} {x : X} (hf : ContinuousAt f x) - {g : X → ∀ j : Fin n, π (i.succAbove j)} (hg : ContinuousAt g x) : +@[deprecated (since := "2025-01-02")] +alias Filter.Tendsto.fin_insertNth := Filter.Tendsto.finInsertNth + +theorem ContinuousAt.finInsertNth + (i : Fin (n + 1)) {f : X → π i} {g : X → ∀ j : Fin n, π (i.succAbove j)} {x : X} + (hf : ContinuousAt f x) (hg : ContinuousAt g x) : ContinuousAt (fun a => i.insertNth (f a) (g a)) x := - hf.tendsto.fin_insertNth i hg + hf.tendsto.finInsertNth i hg + +@[deprecated (since := "2025-01-02")] +alias ContinuousAt.fin_insertNth := ContinuousAt.finInsertNth + +theorem Continuous.finInsertNth + (i : Fin (n + 1)) {f : X → π i} {g : X → ∀ j : Fin n, π (i.succAbove j)} + (hf : Continuous f) (hg : Continuous g) : Continuous fun a => i.insertNth (f a) (g a) := + continuous_iff_continuousAt.2 fun _ => hf.continuousAt.finInsertNth i hg.continuousAt + +@[deprecated (since := "2025-01-02")] +alias Continuous.fin_insertNth := Continuous.finInsertNth -theorem Continuous.fin_insertNth {n} {π : Fin (n + 1) → Type*} [∀ i, TopologicalSpace (π i)] - (i : Fin (n + 1)) {f : X → π i} (hf : Continuous f) {g : X → ∀ j : Fin n, π (i.succAbove j)} - (hg : Continuous g) : Continuous fun a => i.insertNth (f a) (g a) := - continuous_iff_continuousAt.2 fun _ => hf.continuousAt.fin_insertNth i hg.continuousAt +end Fin theorem isOpen_set_pi {i : Set ι} {s : ∀ a, Set (π a)} (hi : i.Finite) (hs : ∀ a ∈ i, IsOpen (s a)) : IsOpen (pi i s) := by diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index a919902aa224f..745a728ee1127 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -1201,17 +1201,61 @@ protected theorem ContinuousOn.iterate {f : α → α} {s : Set α} (hcont : Con | 0 => continuousOn_id | (n + 1) => (hcont.iterate hmaps n).comp hcont hmaps -theorem ContinuousWithinAt.fin_insertNth {n} {π : Fin (n + 1) → Type*} - [∀ i, TopologicalSpace (π i)] (i : Fin (n + 1)) {f : α → π i} {a : α} {s : Set α} - (hf : ContinuousWithinAt f s a) {g : α → ∀ j : Fin n, π (i.succAbove j)} - (hg : ContinuousWithinAt g s a) : ContinuousWithinAt (fun a => i.insertNth (f a) (g a)) s a := - hf.tendsto.fin_insertNth i hg - -nonrec theorem ContinuousOn.fin_insertNth {n} {π : Fin (n + 1) → Type*} - [∀ i, TopologicalSpace (π i)] (i : Fin (n + 1)) {f : α → π i} {s : Set α} - (hf : ContinuousOn f s) {g : α → ∀ j : Fin n, π (i.succAbove j)} (hg : ContinuousOn g s) : +section Fin +variable {n : ℕ} {π : Fin (n + 1) → Type*} [∀ i, TopologicalSpace (π i)] + +theorem ContinuousWithinAt.finCons + {f : α → π 0} {g : α → ∀ j : Fin n, π (Fin.succ j)} {a : α} {s : Set α} + (hf : ContinuousWithinAt f s a) (hg : ContinuousWithinAt g s a) : + ContinuousWithinAt (fun a => Fin.cons (f a) (g a)) s a := + hf.tendsto.finCons hg + +theorem ContinuousOn.finCons {f : α → π 0} {s : Set α} {g : α → ∀ j : Fin n, π (Fin.succ j)} + (hf : ContinuousOn f s) (hg : ContinuousOn g s) : + ContinuousOn (fun a => Fin.cons (f a) (g a)) s := fun a ha => + (hf a ha).finCons (hg a ha) + +theorem ContinuousWithinAt.matrixVecCons {f : α → β} {g : α → Fin n → β} {a : α} {s : Set α} + (hf : ContinuousWithinAt f s a) (hg : ContinuousWithinAt g s a) : + ContinuousWithinAt (fun a => Matrix.vecCons (f a) (g a)) s a := + hf.tendsto.matrixVecCons hg + +theorem ContinuousOn.matrixVecCons {f : α → β} {g : α → Fin n → β} {s : Set α} + (hf : ContinuousOn f s) (hg : ContinuousOn g s) : + ContinuousOn (fun a => Matrix.vecCons (f a) (g a)) s := fun a ha => + (hf a ha).matrixVecCons (hg a ha) + +theorem ContinuousWithinAt.finSnoc + {f : α → ∀ j : Fin n, π (Fin.castSucc j)} {g : α → π (Fin.last _)} {a : α} {s : Set α} + (hf : ContinuousWithinAt f s a) (hg : ContinuousWithinAt g s a) : + ContinuousWithinAt (fun a => Fin.snoc (f a) (g a)) s a := + hf.tendsto.finSnoc hg + +theorem ContinuousOn.finSnoc + {f : α → ∀ j : Fin n, π (Fin.castSucc j)} {g : α → π (Fin.last _)} {s : Set α} + (hf : ContinuousOn f s) (hg : ContinuousOn g s) : + ContinuousOn (fun a => Fin.snoc (f a) (g a)) s := fun a ha => + (hf a ha).finSnoc (hg a ha) + +theorem ContinuousWithinAt.finInsertNth + (i : Fin (n + 1)) {f : α → π i} {g : α → ∀ j : Fin n, π (i.succAbove j)} {a : α} {s : Set α} + (hf : ContinuousWithinAt f s a) (hg : ContinuousWithinAt g s a) : + ContinuousWithinAt (fun a => i.insertNth (f a) (g a)) s a := + hf.tendsto.finInsertNth i hg + +@[deprecated (since := "2025-01-02")] +alias ContinuousWithinAt.fin_insertNth := ContinuousWithinAt.finInsertNth + +theorem ContinuousOn.finInsertNth + (i : Fin (n + 1)) {f : α → π i} {g : α → ∀ j : Fin n, π (i.succAbove j)} {s : Set α} + (hf : ContinuousOn f s) (hg : ContinuousOn g s) : ContinuousOn (fun a => i.insertNth (f a) (g a)) s := fun a ha => - (hf a ha).fin_insertNth i (hg a ha) + (hf a ha).finInsertNth i (hg a ha) + +@[deprecated (since := "2025-01-02")] +alias ContinuousOn.fin_insertNth := ContinuousOn.finInsertNth + +end Fin theorem Set.LeftInvOn.map_nhdsWithin_eq {f : α → β} {g : β → α} {x : β} {s : Set β} (h : LeftInvOn f g s) (hx : f (g x) = x) (hf : ContinuousWithinAt f (g '' s) (g x)) From 06deeaef8d329cdc75e5648214137345f3746b34 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Thu, 2 Jan 2025 12:06:46 +0000 Subject: [PATCH 085/189] feat(Analysis/Normed/Field/Basic): NormOneClass for SubringClass (#20373) --- Mathlib/Analysis/Normed/Field/Basic.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index bd7716e46f8c9..723a44e211ecc 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -1098,6 +1098,10 @@ instance toSeminormedCommRing [SeminormedCommRing R] [_h : SubringClass S R] (s instance toNormedCommRing [NormedCommRing R] [SubringClass S R] (s : S) : NormedCommRing s := { SubringClass.toNormedRing s with mul_comm := mul_comm } +instance toNormOneClass [SeminormedRing R] [NormOneClass R] [SubringClass S R] (s : S) : + NormOneClass s := + .induced s R <| SubringClass.subtype _ + end SubringClass namespace SubfieldClass From 041d59b4c1ad8df53c9ea0224b076f0f3c60ddc4 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 2 Jan 2025 12:06:47 +0000 Subject: [PATCH 086/189] chore(FreeAbelianGroup): don't `dsimp [OfNat.ofNat]` (#20387) There is no need to do this, as there are better lemmas available (at least, after reordering slightly). This also generalizes the `of_one` and `one_def` lemmas to require `One` and not `Monoid`. --- Mathlib/GroupTheory/FreeAbelianGroup.lean | 34 +++++++++++------------ Mathlib/NumberTheory/ADEInequality.lean | 2 -- 2 files changed, 17 insertions(+), 19 deletions(-) diff --git a/Mathlib/GroupTheory/FreeAbelianGroup.lean b/Mathlib/GroupTheory/FreeAbelianGroup.lean index ad652f38aa3fc..41d322ed1f481 100644 --- a/Mathlib/GroupTheory/FreeAbelianGroup.lean +++ b/Mathlib/GroupTheory/FreeAbelianGroup.lean @@ -402,9 +402,20 @@ instance nonUnitalNonAssocRing : NonUnitalNonAssocRing (FreeAbelianGroup α) := end Mul -instance one [One α] : One (FreeAbelianGroup α) := +section One +variable [One α] + +instance one : One (FreeAbelianGroup α) := ⟨of 1⟩ +theorem one_def : (1 : FreeAbelianGroup α) = of 1 := + rfl + +theorem of_one : (of 1 : FreeAbelianGroup α) = 1 := + rfl + +end One + instance nonUnitalRing [Semigroup α] : NonUnitalRing (FreeAbelianGroup α) := { FreeAbelianGroup.nonUnitalNonAssocRing with mul_assoc := fun x y z ↦ by @@ -430,21 +441,16 @@ instance ring : Ring (FreeAbelianGroup α) := { FreeAbelianGroup.nonUnitalRing _, FreeAbelianGroup.one _ with mul_one := fun x ↦ by - dsimp only [(· * ·), Mul.mul, OfNat.ofNat, One.one] - rw [lift.of] + rw [mul_def, one_def, lift.of] refine FreeAbelianGroup.induction_on x rfl (fun L ↦ ?_) (fun L ih ↦ ?_) fun x1 x2 ih1 ih2 ↦ ?_ - · rw [lift.of] - congr 1 - exact mul_one L + · rw [lift.of, mul_one] · rw [map_neg, ih] · rw [map_add, ih1, ih2] one_mul := fun x ↦ by - dsimp only [(· * ·), Mul.mul, OfNat.ofNat, One.one] + simp_rw [mul_def, one_def, lift.of] refine FreeAbelianGroup.induction_on x rfl ?_ ?_ ?_ · intro L - rw [lift.of, lift.of] - congr 1 - exact one_mul L + rw [lift.of, one_mul] · intro L ih rw [map_neg, ih] · intro x1 x2 ih1 ih2 @@ -455,7 +461,7 @@ variable {α} /-- `FreeAbelianGroup.of` is a `MonoidHom` when `α` is a `Monoid`. -/ def ofMulHom : α →* FreeAbelianGroup α where toFun := of - map_one' := rfl + map_one' := of_one _ map_mul' := of_mul @[simp] @@ -505,12 +511,6 @@ theorem liftMonoid_symm_coe (f : FreeAbelianGroup α →+* R) : ⇑(liftMonoid.symm f) = lift.symm (↑f : FreeAbelianGroup α →+ R) := rfl -theorem one_def : (1 : FreeAbelianGroup α) = of 1 := - rfl - -theorem of_one : (of 1 : FreeAbelianGroup α) = 1 := - rfl - end Monoid instance [CommMonoid α] : CommRing (FreeAbelianGroup α) := diff --git a/Mathlib/NumberTheory/ADEInequality.lean b/Mathlib/NumberTheory/ADEInequality.lean index f6cf4a0e5c35e..8d066351a45ad 100644 --- a/Mathlib/NumberTheory/ADEInequality.lean +++ b/Mathlib/NumberTheory/ADEInequality.lean @@ -142,11 +142,9 @@ theorem Admissible.one_lt_sumInv {pqr : Multiset ℕ+} : Admissible pqr → 1 < simp only [lt_add_iff_pos_right, PNat.one_coe, inv_one, Nat.cast_one] apply add_pos <;> simp only [PNat.pos, Nat.cast_pos, inv_pos] · rw [← H, D', sumInv_pqr] - conv_rhs => simp only [OfNat.ofNat, PNat.mk_coe] norm_num all_goals rw [← H, E', sumInv_pqr] - conv_rhs => simp only [OfNat.ofNat, PNat.mk_coe] norm_num theorem lt_three {p q r : ℕ+} (hpq : p ≤ q) (hqr : q ≤ r) (H : 1 < sumInv {p, q, r}) : p < 3 := by From 6d8ce693e6385b761fab6c2f523288a776e2fa51 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Thu, 2 Jan 2025 12:45:36 +0000 Subject: [PATCH 087/189] chore(technical-debt-metrics): Add ofNat to technical debt metrics (#20394) --- scripts/technical-debt-metrics.sh | 1 + 1 file changed, 1 insertion(+) diff --git a/scripts/technical-debt-metrics.sh b/scripts/technical-debt-metrics.sh index b3f3df14ec9aa..7c982cadbcb20 100755 --- a/scripts/technical-debt-metrics.sh +++ b/scripts/technical-debt-metrics.sh @@ -80,6 +80,7 @@ titlesPathsAndRegexes=( "adaptation notes" "*" "adaptation_note" "disabled simpNF lints" "*" "nolint simpNF" "erw" "*" "erw \[" + "ofNat" "*" "See note \[no_index around OfNat.ofNat\]" "maxHeartBeats modifications" ":^MathlibTest" "^ *set_option .*maxHeartbeats" ) From df255e235cde3ac9e4325de51cfa2bc82b7b2416 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 2 Jan 2025 13:01:02 +0000 Subject: [PATCH 088/189] chore(RingTheory/Artinian): Split out `Artinian/Ring.lean` (#20350) Co-authored-by: Andrew Yang Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 4 +- Mathlib/Algebra/Lie/Derivation/Basic.lean | 1 + Mathlib/Algebra/Lie/Nilpotent.lean | 3 +- Mathlib/Algebra/Lie/Quotient.lean | 1 + Mathlib/Algebra/Lie/Subalgebra.lean | 2 +- Mathlib/Algebra/Lie/Submodule.lean | 3 +- Mathlib/Algebra/Lie/TensorProduct.lean | 1 + Mathlib/Algebra/Lie/Weights/Basic.lean | 2 +- Mathlib/LinearAlgebra/Eigenspace/Zero.lean | 3 +- Mathlib/LinearAlgebra/Matrix/ZPow.lean | 1 - Mathlib/LinearAlgebra/Semisimple.lean | 5 +- Mathlib/RingTheory/Artinian/Instances.lean | 31 ++++ .../{Artinian.lean => Artinian/Module.lean} | 133 ++++-------------- Mathlib/RingTheory/Artinian/Ring.lean | 116 +++++++++++++++ Mathlib/RingTheory/FiniteLength.lean | 3 +- Mathlib/RingTheory/Jacobson/Ring.lean | 2 +- Mathlib/RingTheory/Noetherian/Defs.lean | 28 ++++ Mathlib/RingTheory/Noetherian/Orzech.lean | 25 ---- .../RingTheory/Regular/RegularSequence.lean | 4 +- Mathlib/RingTheory/Unramified/Field.lean | 2 +- 20 files changed, 228 insertions(+), 142 deletions(-) create mode 100644 Mathlib/RingTheory/Artinian/Instances.lean rename Mathlib/RingTheory/{Artinian.lean => Artinian/Module.lean} (81%) create mode 100644 Mathlib/RingTheory/Artinian/Ring.lean diff --git a/Mathlib.lean b/Mathlib.lean index c15d73c88f662..6740698922ff9 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4327,7 +4327,9 @@ import Mathlib.RingTheory.AlgebraicIndependent.Defs import Mathlib.RingTheory.AlgebraicIndependent.RankAndCardinality import Mathlib.RingTheory.AlgebraicIndependent.TranscendenceBasis import Mathlib.RingTheory.AlgebraicIndependent.Transcendental -import Mathlib.RingTheory.Artinian +import Mathlib.RingTheory.Artinian.Instances +import Mathlib.RingTheory.Artinian.Module +import Mathlib.RingTheory.Artinian.Ring import Mathlib.RingTheory.Bezout import Mathlib.RingTheory.Bialgebra.Basic import Mathlib.RingTheory.Bialgebra.Equiv diff --git a/Mathlib/Algebra/Lie/Derivation/Basic.lean b/Mathlib/Algebra/Lie/Derivation/Basic.lean index 306dd4f098b2f..09dfb6b3c58cc 100644 --- a/Mathlib/Algebra/Lie/Derivation/Basic.lean +++ b/Mathlib/Algebra/Lie/Derivation/Basic.lean @@ -6,6 +6,7 @@ Authors: Frédéric Marbach import Mathlib.Algebra.Lie.Basic import Mathlib.Algebra.Lie.OfAssociative import Mathlib.Algebra.Lie.Subalgebra +import Mathlib.RingTheory.Noetherian.Basic /-! # Lie derivations diff --git a/Mathlib/Algebra/Lie/Nilpotent.lean b/Mathlib/Algebra/Lie/Nilpotent.lean index 87bf4540070a5..7abb7b337eb4f 100644 --- a/Mathlib/Algebra/Lie/Nilpotent.lean +++ b/Mathlib/Algebra/Lie/Nilpotent.lean @@ -7,9 +7,10 @@ import Mathlib.Algebra.Lie.BaseChange import Mathlib.Algebra.Lie.Solvable import Mathlib.Algebra.Lie.Quotient import Mathlib.Algebra.Lie.Normalizer +import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.LinearAlgebra.Eigenspace.Basic import Mathlib.Order.Filter.AtTopBot -import Mathlib.RingTheory.Artinian +import Mathlib.RingTheory.Artinian.Module import Mathlib.RingTheory.Nilpotent.Lemmas /-! diff --git a/Mathlib/Algebra/Lie/Quotient.lean b/Mathlib/Algebra/Lie/Quotient.lean index 367d16adeac24..7a8c32ffb48c3 100644 --- a/Mathlib/Algebra/Lie/Quotient.lean +++ b/Mathlib/Algebra/Lie/Quotient.lean @@ -6,6 +6,7 @@ Authors: Oliver Nash import Mathlib.Algebra.Lie.Submodule import Mathlib.Algebra.Lie.OfAssociative import Mathlib.LinearAlgebra.Isomorphisms +import Mathlib.RingTheory.Noetherian.Basic /-! # Quotients of Lie algebras and Lie modules diff --git a/Mathlib/Algebra/Lie/Subalgebra.lean b/Mathlib/Algebra/Lie/Subalgebra.lean index 661062dd102f7..5cf5fab92ae25 100644 --- a/Mathlib/Algebra/Lie/Subalgebra.lean +++ b/Mathlib/Algebra/Lie/Subalgebra.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ import Mathlib.Algebra.Lie.Basic -import Mathlib.RingTheory.Artinian +import Mathlib.RingTheory.Artinian.Module /-! # Lie subalgebras diff --git a/Mathlib/Algebra/Lie/Submodule.lean b/Mathlib/Algebra/Lie/Submodule.lean index 24345ff122879..5b057903a082c 100644 --- a/Mathlib/Algebra/Lie/Submodule.lean +++ b/Mathlib/Algebra/Lie/Submodule.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ import Mathlib.Algebra.Lie.Subalgebra -import Mathlib.RingTheory.Artinian -import Mathlib.RingTheory.Noetherian.Orzech +import Mathlib.LinearAlgebra.Finsupp.Span /-! # Lie submodules of a Lie algebra diff --git a/Mathlib/Algebra/Lie/TensorProduct.lean b/Mathlib/Algebra/Lie/TensorProduct.lean index e942981a6181a..240be784917b3 100644 --- a/Mathlib/Algebra/Lie/TensorProduct.lean +++ b/Mathlib/Algebra/Lie/TensorProduct.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ import Mathlib.Algebra.Lie.Abelian +import Mathlib.LinearAlgebra.TensorProduct.Tower /-! # Tensor products of Lie modules diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index 0e9698e5fa0d0..d6c6f9d9c417c 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.Ring.Divisibility.Lemmas import Mathlib.Algebra.Lie.Nilpotent import Mathlib.Algebra.Lie.Engel import Mathlib.LinearAlgebra.Eigenspace.Pi -import Mathlib.RingTheory.Artinian +import Mathlib.RingTheory.Artinian.Module import Mathlib.LinearAlgebra.Trace import Mathlib.LinearAlgebra.FreeModule.PID diff --git a/Mathlib/LinearAlgebra/Eigenspace/Zero.lean b/Mathlib/LinearAlgebra/Eigenspace/Zero.lean index 124109d510357..d3cffd7da595c 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Zero.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Zero.lean @@ -3,11 +3,12 @@ Copyright (c) 2024 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ +import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.LinearAlgebra.Charpoly.ToMatrix import Mathlib.LinearAlgebra.Determinant import Mathlib.LinearAlgebra.Eigenspace.Minpoly import Mathlib.LinearAlgebra.FreeModule.StrongRankCondition -import Mathlib.RingTheory.Artinian +import Mathlib.RingTheory.Artinian.Module /-! # Results on the eigenvalue 0 diff --git a/Mathlib/LinearAlgebra/Matrix/ZPow.lean b/Mathlib/LinearAlgebra/Matrix/ZPow.lean index dd11fe3101db7..3535d8a612953 100644 --- a/Mathlib/LinearAlgebra/Matrix/ZPow.lean +++ b/Mathlib/LinearAlgebra/Matrix/ZPow.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Yakov Pechersky. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yakov Pechersky -/ -import Mathlib.Data.Int.Bitwise import Mathlib.LinearAlgebra.Matrix.NonsingularInverse import Mathlib.LinearAlgebra.Matrix.Symmetric diff --git a/Mathlib/LinearAlgebra/Semisimple.lean b/Mathlib/LinearAlgebra/Semisimple.lean index f96061725cfc2..c7032628a6166 100644 --- a/Mathlib/LinearAlgebra/Semisimple.lean +++ b/Mathlib/LinearAlgebra/Semisimple.lean @@ -3,13 +3,12 @@ Copyright (c) 2024 Oliver Nash. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ +import Mathlib.Algebra.Module.Torsion import Mathlib.FieldTheory.Perfect import Mathlib.LinearAlgebra.AnnihilatingPolynomial -import Mathlib.LinearAlgebra.Basis.VectorSpace -import Mathlib.RingTheory.Artinian +import Mathlib.RingTheory.Artinian.Instances import Mathlib.RingTheory.Ideal.Quotient.Nilpotent import Mathlib.RingTheory.SimpleModule -import Mathlib.Algebra.Module.Torsion /-! # Semisimple linear endomorphisms diff --git a/Mathlib/RingTheory/Artinian/Instances.lean b/Mathlib/RingTheory/Artinian/Instances.lean new file mode 100644 index 0000000000000..a5b60aa963ede --- /dev/null +++ b/Mathlib/RingTheory/Artinian/Instances.lean @@ -0,0 +1,31 @@ +/- +Copyright (c) 2024 Junyan Xu. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Junyan Xu +-/ +import Mathlib.Algebra.Divisibility.Prod +import Mathlib.Algebra.Polynomial.FieldDivision +import Mathlib.RingTheory.Artinian.Module +import Mathlib.RingTheory.SimpleModule + +/-! +# Instances related to Artinian rings + +We show that every reduced Artinian ring and the polynomial ring over it +are decomposition monoids, and every reduced Artinian ring is semisimple. +-/ + +namespace IsArtinianRing + +variable (R : Type*) [CommRing R] [IsArtinianRing R] [IsReduced R] + +attribute [local instance] subtype_isMaximal_finite fieldOfSubtypeIsMaximal + +instance : DecompositionMonoid R := MulEquiv.decompositionMonoid (equivPi R) + +instance : DecompositionMonoid (Polynomial R) := + MulEquiv.decompositionMonoid <| (Polynomial.mapEquiv <| equivPi R).trans (Polynomial.piEquiv _) + +theorem isSemisimpleRing_of_isReduced : IsSemisimpleRing R := (equivPi R).symm.isSemisimpleRing + +end IsArtinianRing diff --git a/Mathlib/RingTheory/Artinian.lean b/Mathlib/RingTheory/Artinian/Module.lean similarity index 81% rename from Mathlib/RingTheory/Artinian.lean rename to Mathlib/RingTheory/Artinian/Module.lean index 6bf257f5cca5f..53f91c83eee98 100644 --- a/Mathlib/RingTheory/Artinian.lean +++ b/Mathlib/RingTheory/Artinian/Module.lean @@ -3,16 +3,11 @@ Copyright (c) 2021 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.Divisibility.Prod -import Mathlib.Algebra.Polynomial.FieldDivision import Mathlib.Data.SetLike.Fintype import Mathlib.Order.Filter.EventuallyConst -import Mathlib.RingTheory.Nakayama -import Mathlib.RingTheory.SimpleModule -import Mathlib.Tactic.RSuffices -import Mathlib.Tactic.StacksAttribute -import Mathlib.RingTheory.LocalRing.Basic +import Mathlib.RingTheory.Ideal.Quotient.Operations import Mathlib.RingTheory.Nilpotent.Lemmas +import Mathlib.RingTheory.Noetherian.Defs /-! # Artinian rings and modules @@ -34,12 +29,6 @@ Let `R` be a ring and let `M` and `P` be `R`-modules. Let `N` be an `R`-submodul ## Main results -* `IsArtinianRing.localization_surjective`: the canonical homomorphism from a commutative artinian - ring to any localization of itself is surjective. - -* `IsArtinianRing.isNilpotent_jacobson_bot`: the Jacobson radical of a commutative artinian ring - is a nilpotent ideal. (TODO: generalize to noncommutative rings.) - * `IsArtinianRing.primeSpectrum_finite`, `IsArtinianRing.isMaximal_of_isPrime`: there are only finitely prime ideals in a commutative artinian ring, and each of them is maximal. @@ -47,6 +36,7 @@ Let `R` be a ring and let `M` and `P` be `R`-modules. Let `N` be an `R`-submodul product of fields (and therefore is a semisimple ring and a decomposition monoid; moreover `R[X]` is also a decomposition monoid). + ## References * [M. F. Atiyah and I. G. Macdonald, *Introduction to commutative algebra*][atiyah-macdonald] @@ -346,10 +336,25 @@ end IsArtinian end CommRing +theorem isArtinian_of_submodule_of_artinian (R M) [Ring R] [AddCommGroup M] [Module R M] + (N : Submodule R M) (_ : IsArtinian R M) : IsArtinian R N := inferInstance + +/-- If `M / S / R` is a scalar tower, and `M / R` is Artinian, then `M / S` is also Artinian. -/ +theorem isArtinian_of_tower (R) {S M} [CommRing R] [Ring S] [AddCommGroup M] [Algebra R S] + [Module S M] [Module R M] [IsScalarTower R S M] (h : IsArtinian R M) : IsArtinian S M := + ⟨(Submodule.restrictScalarsEmbedding R S M).wellFounded h.wf⟩ + +-- See `Mathlib.RingTheory.Artinian.Ring` +assert_not_exists IsLocalization +assert_not_exists LocalRing + /-- A ring is Artinian if it is Artinian as a module over itself. Strictly speaking, this should be called `IsLeftArtinianRing` but we omit the `Left` for -convenience in the commutative case. For a right Artinian ring, use `IsArtinian Rᵐᵒᵖ R`. -/ +convenience in the commutative case. For a right Artinian ring, use `IsArtinian Rᵐᵒᵖ R`. + +For equivalent definitions, see `Mathlib.RingTheory.Artinian.Ring`. +-/ @[stacks 00J5] abbrev IsArtinianRing (R) [Ring R] := IsArtinian R R @@ -363,14 +368,6 @@ theorem Ring.isArtinian_of_zero_eq_one {R} [Ring R] (h01 : (0 : R) = 1) : IsArti have := subsingleton_of_zero_eq_one h01 inferInstance -theorem isArtinian_of_submodule_of_artinian (R M) [Ring R] [AddCommGroup M] [Module R M] - (N : Submodule R M) (_ : IsArtinian R M) : IsArtinian R N := inferInstance - -/-- If `M / S / R` is a scalar tower, and `M / R` is Artinian, then `M / S` is also Artinian. -/ -theorem isArtinian_of_tower (R) {S M} [CommRing R] [Ring S] [AddCommGroup M] [Algebra R S] - [Module S M] [Module R M] [IsScalarTower R S M] (h : IsArtinian R M) : IsArtinian S M := - ⟨(Submodule.restrictScalarsEmbedding R S M).wellFounded h.wf⟩ - instance (R) [CommRing R] [IsArtinianRing R] (I : Ideal R) : IsArtinianRing (R ⧸ I) := isArtinian_of_tower R inferInstance @@ -418,81 +415,21 @@ instance isArtinianRing_range {R} [Ring R] {S} [Ring S] (f : R →+* S) [IsArtin namespace IsArtinianRing -open IsArtinian - variable {R : Type*} [CommRing R] [IsArtinianRing R] -@[stacks 00J8] -theorem isNilpotent_jacobson_bot : IsNilpotent (Ideal.jacobson (⊥ : Ideal R)) := by - let Jac := Ideal.jacobson (⊥ : Ideal R) - let f : ℕ →o (Ideal R)ᵒᵈ := ⟨fun n => Jac ^ n, fun _ _ h => Ideal.pow_le_pow_right h⟩ - obtain ⟨n, hn⟩ : ∃ n, ∀ m, n ≤ m → Jac ^ n = Jac ^ m := IsArtinian.monotone_stabilizes f - refine ⟨n, ?_⟩ - let J : Ideal R := annihilator (Jac ^ n) - suffices J = ⊤ by - have hJ : J • Jac ^ n = ⊥ := annihilator_smul (Jac ^ n) - simpa only [this, top_smul, Ideal.zero_eq_bot] using hJ - by_contra hJ - change J ≠ ⊤ at hJ - rcases IsArtinian.set_has_minimal { J' : Ideal R | J < J' } ⟨⊤, hJ.lt_top⟩ with - ⟨J', hJJ' : J < J', hJ' : ∀ I, J < I → ¬I < J'⟩ - rcases SetLike.exists_of_lt hJJ' with ⟨x, hxJ', hxJ⟩ - obtain rfl : J ⊔ Ideal.span {x} = J' := by - apply eq_of_le_of_not_lt _ (hJ' (J ⊔ Ideal.span {x}) _) - · exact sup_le hJJ'.le (span_le.2 (singleton_subset_iff.2 hxJ')) - · rw [SetLike.lt_iff_le_and_exists] - exact ⟨le_sup_left, ⟨x, mem_sup_right (mem_span_singleton_self x), hxJ⟩⟩ - have : J ⊔ Jac • Ideal.span {x} ≤ J ⊔ Ideal.span {x} := - sup_le_sup_left (smul_le.2 fun _ _ _ => Submodule.smul_mem _ _) _ - have : Jac * Ideal.span {x} ≤ J := by -- Need version 4 of Nakayama's lemma on Stacks - by_contra H - refine H (Ideal.mul_le_left.trans (le_of_le_smul_of_le_jacobson_bot (fg_span_singleton _) le_rfl - (le_sup_right.trans_eq (this.eq_of_not_lt (hJ' _ ?_)).symm))) - exact lt_of_le_of_ne le_sup_left fun h => H <| h.symm ▸ le_sup_right - have : Ideal.span {x} * Jac ^ (n + 1) ≤ ⊥ := calc - Ideal.span {x} * Jac ^ (n + 1) = Ideal.span {x} * Jac * Jac ^ n := by - rw [pow_succ', ← mul_assoc] - _ ≤ J * Jac ^ n := mul_le_mul (by rwa [mul_comm]) le_rfl - _ = ⊥ := by simp [J] - refine hxJ (mem_annihilator.2 fun y hy => (mem_bot R).1 ?_) - refine this (mul_mem_mul (mem_span_singleton_self x) ?_) - rwa [← hn (n + 1) (Nat.le_succ _)] - -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 - intro r' - obtain ⟨r₁, s, rfl⟩ := IsLocalization.mk'_surjective S r' - -- TODO: can `rsuffices` be used to move the `exact` below before the proof of this `obtain`? - obtain ⟨r₂, h⟩ : ∃ r : R, IsLocalization.mk' L 1 s = algebraMap R L r := by - obtain ⟨n, r, hr⟩ := IsArtinian.exists_pow_succ_smul_dvd (s : R) (1 : R) - use r - rw [smul_eq_mul, smul_eq_mul, pow_succ, mul_assoc] at hr - apply_fun algebraMap R L at hr - simp only [map_mul] at hr - rw [← IsLocalization.mk'_one (M := S) L, IsLocalization.mk'_eq_iff_eq, mul_one, - Submonoid.coe_one, ← (IsLocalization.map_units L (s ^ n)).mul_left_cancel hr, map_mul] - exact ⟨r₁ * r₂, by rw [IsLocalization.mk'_eq_mul_mk'_one, map_mul, h]⟩ - -theorem localization_artinian : IsArtinianRing L := - (localization_surjective S L).isArtinianRing - -/-- `IsArtinianRing.localization_artinian` can't be made an instance, as it would make `S` + `R` -into metavariables. However, this is safe. -/ -instance : IsArtinianRing (Localization S) := - localization_artinian S _ - -end Localization +variable (R) in +lemma isField_of_isDomain [IsDomain R] : IsField R := by + refine ⟨Nontrivial.exists_pair_ne, mul_comm, fun {x} hx ↦ ?_⟩ + obtain ⟨n, y, hy⟩ := IsArtinian.exists_pow_succ_smul_dvd x (1 : R) + replace hy : x ^ n * (x * y - 1) = 0 := by + rw [mul_sub, sub_eq_zero] + convert hy using 1 + simp [Nat.succ_eq_add_one, pow_add, mul_assoc] + rw [mul_eq_zero, sub_eq_zero] at hy + exact ⟨_, hy.resolve_left <| pow_ne_zero _ hx⟩ instance isMaximal_of_isPrime (p : Ideal R) [p.IsPrime] : p.IsMaximal := - Ideal.Quotient.maximal_of_isField _ <| - (MulEquiv.ofBijective _ ⟨IsFractionRing.injective (R ⧸ p) (FractionRing (R ⧸ p)), - localization_surjective (nonZeroDivisors (R ⧸ p)) (FractionRing (R ⧸ p))⟩).isField _ - <| Field.toIsField <| FractionRing (R ⧸ p) + Ideal.Quotient.maximal_of_isField _ (isField_of_isDomain _) lemma isPrime_iff_isMaximal (p : Ideal R) : p.IsPrime ↔ p.IsMaximal := ⟨fun _ ↦ isMaximal_of_isPrime p, fun h ↦ h.isPrime⟩ @@ -500,7 +437,7 @@ lemma isPrime_iff_isMaximal (p : Ideal R) : p.IsPrime ↔ p.IsMaximal := variable (R) in lemma primeSpectrum_finite : {I : Ideal R | I.IsPrime}.Finite := by set Spec := {I : Ideal R | I.IsPrime} - obtain ⟨_, ⟨s, rfl⟩, H⟩ := set_has_minimal + obtain ⟨_, ⟨s, rfl⟩, H⟩ := IsArtinian.set_has_minimal (range (Finset.inf · Subtype.val : Finset Spec → Ideal R)) ⟨⊤, ∅, by simp⟩ refine Set.finite_def.2 ⟨s, fun p ↦ ?_⟩ classical @@ -537,12 +474,4 @@ noncomputable def equivPi [IsReduced R] : R ≃+* ∀ I : {I : Ideal R | I.IsMax .trans (.symm <| .quotientBot R) <| .trans (Ideal.quotEquivOfEq (nilradical_eq_zero R).symm) (quotNilradicalEquivPi R) -instance [IsReduced R] : DecompositionMonoid R := MulEquiv.decompositionMonoid (equivPi R) - -instance [IsReduced R] : DecompositionMonoid (Polynomial R) := - MulEquiv.decompositionMonoid <| (Polynomial.mapEquiv <| equivPi R).trans (Polynomial.piEquiv _) - -theorem isSemisimpleRing_of_isReduced [IsReduced R] : IsSemisimpleRing R := - (equivPi R).symm.isSemisimpleRing - end IsArtinianRing diff --git a/Mathlib/RingTheory/Artinian/Ring.lean b/Mathlib/RingTheory/Artinian/Ring.lean new file mode 100644 index 0000000000000..3c0e9a2fde380 --- /dev/null +++ b/Mathlib/RingTheory/Artinian/Ring.lean @@ -0,0 +1,116 @@ +/- +Copyright (c) 2021 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Junyan Xu, Jujian Zhang +-/ +import Mathlib.RingTheory.Artinian.Module +import Mathlib.RingTheory.Localization.Defs +import Mathlib.RingTheory.Nakayama + +/-! +# Artinian rings + +A ring is said to be left (or right) Artinian if it is Artinian as a left (or right) module over +itself, or simply Artinian if it is both left and right Artinian. + +## Main definitions + +* `IsArtinianRing R` is the proposition that `R` is a left Artinian ring. + +## Main results + +* `IsArtinianRing.localization_surjective`: the canonical homomorphism from a commutative artinian + ring to any localization of itself is surjective. + +* `IsArtinianRing.isNilpotent_jacobson_bot`: the Jacobson radical of a commutative artinian ring + is a nilpotent ideal. (TODO: generalize to noncommutative rings.) + +## Implementation Details + +The predicate `IsArtinianRing` is defined in `Mathlib.RingTheory.Artinian.Ring` instead, so that we +can apply basic API on artinian modules to division rings without a heavy import. + +## References + +* [M. F. Atiyah and I. G. Macdonald, *Introduction to commutative algebra*][atiyah-macdonald] +* [samuel] + +## Tags + +Artinian, artinian, Artinian ring, artinian ring + +-/ + +open Set Submodule IsArtinian + +namespace IsArtinianRing + +variable {R : Type*} [CommRing R] [IsArtinianRing R] + +@[stacks 00J8] +theorem isNilpotent_jacobson_bot : IsNilpotent (Ideal.jacobson (⊥ : Ideal R)) := by + let Jac := Ideal.jacobson (⊥ : Ideal R) + let f : ℕ →o (Ideal R)ᵒᵈ := ⟨fun n => Jac ^ n, fun _ _ h => Ideal.pow_le_pow_right h⟩ + obtain ⟨n, hn⟩ : ∃ n, ∀ m, n ≤ m → Jac ^ n = Jac ^ m := IsArtinian.monotone_stabilizes f + refine ⟨n, ?_⟩ + let J : Ideal R := annihilator (Jac ^ n) + suffices J = ⊤ by + have hJ : J • Jac ^ n = ⊥ := annihilator_smul (Jac ^ n) + simpa only [this, top_smul, Ideal.zero_eq_bot] using hJ + by_contra hJ + change J ≠ ⊤ at hJ + rcases IsArtinian.set_has_minimal { J' : Ideal R | J < J' } ⟨⊤, hJ.lt_top⟩ with + ⟨J', hJJ' : J < J', hJ' : ∀ I, J < I → ¬I < J'⟩ + rcases SetLike.exists_of_lt hJJ' with ⟨x, hxJ', hxJ⟩ + obtain rfl : J ⊔ Ideal.span {x} = J' := by + apply eq_of_le_of_not_lt _ (hJ' (J ⊔ Ideal.span {x}) _) + · exact sup_le hJJ'.le (span_le.2 (singleton_subset_iff.2 hxJ')) + · rw [SetLike.lt_iff_le_and_exists] + exact ⟨le_sup_left, ⟨x, mem_sup_right (mem_span_singleton_self x), hxJ⟩⟩ + have : J ⊔ Jac • Ideal.span {x} ≤ J ⊔ Ideal.span {x} := + sup_le_sup_left (smul_le.2 fun _ _ _ => Submodule.smul_mem _ _) _ + have : Jac * Ideal.span {x} ≤ J := by -- Need version 4 of Nakayama's lemma on Stacks + by_contra H + refine H (Ideal.mul_le_left.trans (le_of_le_smul_of_le_jacobson_bot (fg_span_singleton _) le_rfl + (le_sup_right.trans_eq (this.eq_of_not_lt (hJ' _ ?_)).symm))) + exact lt_of_le_of_ne le_sup_left fun h => H <| h.symm ▸ le_sup_right + have : Ideal.span {x} * Jac ^ (n + 1) ≤ ⊥ := calc + Ideal.span {x} * Jac ^ (n + 1) = Ideal.span {x} * Jac * Jac ^ n := by + rw [pow_succ', ← mul_assoc] + _ ≤ J * Jac ^ n := mul_le_mul (by rwa [mul_comm]) le_rfl + _ = ⊥ := by simp [J] + refine hxJ (mem_annihilator.2 fun y hy => (mem_bot R).1 ?_) + refine this (mul_mem_mul (mem_span_singleton_self x) ?_) + rwa [← hn (n + 1) (Nat.le_succ _)] + +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 + intro r' + obtain ⟨r₁, s, rfl⟩ := IsLocalization.mk'_surjective S r' + -- TODO: can `rsuffices` be used to move the `exact` below before the proof of this `obtain`? + obtain ⟨r₂, h⟩ : ∃ r : R, IsLocalization.mk' L 1 s = algebraMap R L r := by + obtain ⟨n, r, hr⟩ := IsArtinian.exists_pow_succ_smul_dvd (s : R) (1 : R) + use r + rw [smul_eq_mul, smul_eq_mul, pow_succ, mul_assoc] at hr + apply_fun algebraMap R L at hr + simp only [map_mul] at hr + rw [← IsLocalization.mk'_one (M := S) L, IsLocalization.mk'_eq_iff_eq, mul_one, + Submonoid.coe_one, ← (IsLocalization.map_units L (s ^ n)).mul_left_cancel hr, map_mul] + exact ⟨r₁ * r₂, by rw [IsLocalization.mk'_eq_mul_mk'_one, map_mul, h]⟩ + +theorem localization_artinian : IsArtinianRing L := + (localization_surjective S L).isArtinianRing + +/-- `IsArtinianRing.localization_artinian` can't be made an instance, as it would make `S` + `R` +into metavariables. However, this is safe. -/ +instance : IsArtinianRing (Localization S) := + localization_artinian S _ + +end Localization + +end IsArtinianRing diff --git a/Mathlib/RingTheory/FiniteLength.lean b/Mathlib/RingTheory/FiniteLength.lean index 8eea41b4e20e2..90ac5c06d9d02 100644 --- a/Mathlib/RingTheory/FiniteLength.lean +++ b/Mathlib/RingTheory/FiniteLength.lean @@ -3,7 +3,8 @@ Copyright (c) 2024 Junyan Xu. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Junyan Xu -/ -import Mathlib.RingTheory.Artinian +import Mathlib.RingTheory.Artinian.Module +import Mathlib.RingTheory.SimpleModule /-! # Modules of finite length diff --git a/Mathlib/RingTheory/Jacobson/Ring.lean b/Mathlib/RingTheory/Jacobson/Ring.lean index c2b35f330796f..68129d02d2fcb 100644 --- a/Mathlib/RingTheory/Jacobson/Ring.lean +++ b/Mathlib/RingTheory/Jacobson/Ring.lean @@ -6,7 +6,7 @@ Authors: Devon Tuma import Mathlib.RingTheory.Localization.Away.Basic import Mathlib.RingTheory.Ideal.Over import Mathlib.RingTheory.Jacobson.Polynomial -import Mathlib.RingTheory.Artinian +import Mathlib.RingTheory.Artinian.Module /-! # Jacobson Rings diff --git a/Mathlib/RingTheory/Noetherian/Defs.lean b/Mathlib/RingTheory/Noetherian/Defs.lean index e14e88932c6a3..2385e93db6919 100644 --- a/Mathlib/RingTheory/Noetherian/Defs.lean +++ b/Mathlib/RingTheory/Noetherian/Defs.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kevin Buzzard -/ import Mathlib.RingTheory.Finiteness.Basic +import Mathlib.Order.Filter.AtTopBot /-! # Noetherian rings and modules @@ -163,6 +164,33 @@ theorem monotone_stabilizes_iff_noetherian : (∀ f : ℕ →o Submodule R M, ∃ n, ∀ m, n ≤ m → f n = f m) ↔ IsNoetherian R M := by rw [isNoetherian_iff, WellFounded.monotone_chain_condition] +variable [IsNoetherian R M] + +open Filter +/-- For an endomorphism of a Noetherian module, any sufficiently large iterate has disjoint kernel +and range. -/ +theorem LinearMap.eventually_disjoint_ker_pow_range_pow (f : M →ₗ[R] M) : + ∀ᶠ n in atTop, Disjoint (LinearMap.ker (f ^ n)) (LinearMap.range (f ^ n)) := by + obtain ⟨n, hn : ∀ m, n ≤ m → LinearMap.ker (f ^ n) = LinearMap.ker (f ^ m)⟩ := + monotone_stabilizes_iff_noetherian.mpr inferInstance f.iterateKer + refine eventually_atTop.mpr ⟨n, fun m hm ↦ disjoint_iff.mpr ?_⟩ + rw [← hn _ hm, Submodule.eq_bot_iff] + rintro - ⟨hx, ⟨x, rfl⟩⟩ + apply LinearMap.pow_map_zero_of_le hm + replace hx : x ∈ LinearMap.ker (f ^ (n + m)) := by + simpa [f.pow_apply n, f.pow_apply m, ← f.pow_apply (n + m), ← iterate_add_apply] using hx + rwa [← hn _ (n.le_add_right m)] at hx + +lemma LinearMap.eventually_iSup_ker_pow_eq (f : M →ₗ[R] M) : + ∀ᶠ n in atTop, ⨆ m, LinearMap.ker (f ^ m) = LinearMap.ker (f ^ n) := by + obtain ⟨n, hn : ∀ m, n ≤ m → ker (f ^ n) = ker (f ^ m)⟩ := + monotone_stabilizes_iff_noetherian.mpr inferInstance f.iterateKer + refine eventually_atTop.mpr ⟨n, fun m hm ↦ ?_⟩ + refine le_antisymm (iSup_le fun l ↦ ?_) (le_iSup (fun i ↦ LinearMap.ker (f ^ i)) m) + rcases le_or_lt m l with h | h + · rw [← hn _ (hm.trans h), hn _ hm] + · exact f.iterateKer.monotone h.le + end /-- A (semi)ring is Noetherian if it is Noetherian as a module over itself, diff --git a/Mathlib/RingTheory/Noetherian/Orzech.lean b/Mathlib/RingTheory/Noetherian/Orzech.lean index da0b31cf04f08..14ba706af1464 100644 --- a/Mathlib/RingTheory/Noetherian/Orzech.lean +++ b/Mathlib/RingTheory/Noetherian/Orzech.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kevin Buzzard -/ import Mathlib.Algebra.Module.Submodule.IterateMapComap -import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Order.PartialSups import Mathlib.RingTheory.Noetherian.Basic import Mathlib.RingTheory.OrzechProperty @@ -33,30 +32,6 @@ universe w variable {R M P : Type*} {N : Type w} [Ring R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] [AddCommGroup P] [Module R P] [IsNoetherian R M] -/-- For an endomorphism of a Noetherian module, any sufficiently large iterate has disjoint kernel -and range. -/ -theorem LinearMap.eventually_disjoint_ker_pow_range_pow (f : M →ₗ[R] M) : - ∀ᶠ n in atTop, Disjoint (LinearMap.ker (f ^ n)) (LinearMap.range (f ^ n)) := by - obtain ⟨n, hn : ∀ m, n ≤ m → LinearMap.ker (f ^ n) = LinearMap.ker (f ^ m)⟩ := - monotone_stabilizes_iff_noetherian.mpr inferInstance f.iterateKer - refine eventually_atTop.mpr ⟨n, fun m hm ↦ disjoint_iff.mpr ?_⟩ - rw [← hn _ hm, Submodule.eq_bot_iff] - rintro - ⟨hx, ⟨x, rfl⟩⟩ - apply LinearMap.pow_map_zero_of_le hm - replace hx : x ∈ LinearMap.ker (f ^ (n + m)) := by - simpa [f.pow_apply n, f.pow_apply m, ← f.pow_apply (n + m), ← iterate_add_apply] using hx - rwa [← hn _ (n.le_add_right m)] at hx - -lemma LinearMap.eventually_iSup_ker_pow_eq (f : M →ₗ[R] M) : - ∀ᶠ n in atTop, ⨆ m, LinearMap.ker (f ^ m) = LinearMap.ker (f ^ n) := by - obtain ⟨n, hn : ∀ m, n ≤ m → ker (f ^ n) = ker (f ^ m)⟩ := - monotone_stabilizes_iff_noetherian.mpr inferInstance f.iterateKer - refine eventually_atTop.mpr ⟨n, fun m hm ↦ ?_⟩ - refine le_antisymm (iSup_le fun l ↦ ?_) (le_iSup (fun i ↦ LinearMap.ker (f ^ i)) m) - rcases le_or_lt m l with h | h - · rw [← hn _ (hm.trans h), hn _ hm] - · exact f.iterateKer.monotone h.le - /-- **Orzech's theorem** for Noetherian modules: if `R` is a ring (not necessarily commutative), `M` and `N` are `R`-modules, `M` is Noetherian, `i : N →ₗ[R] M` is injective, `f : N →ₗ[R] M` is surjective, then `f` is also injective. The proof here is adapted from diff --git a/Mathlib/RingTheory/Regular/RegularSequence.lean b/Mathlib/RingTheory/Regular/RegularSequence.lean index 8beb65b18807c..0b49876fcc391 100644 --- a/Mathlib/RingTheory/Regular/RegularSequence.lean +++ b/Mathlib/RingTheory/Regular/RegularSequence.lean @@ -4,9 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Brendan Murphy -/ import Mathlib.RingTheory.Regular.IsSMulRegular -import Mathlib.RingTheory.Artinian +import Mathlib.RingTheory.Artinian.Module +import Mathlib.RingTheory.Nakayama import Mathlib.Algebra.Equiv.TransferInstance import Mathlib.RingTheory.LocalRing.MaximalIdeal.Basic +import Mathlib.RingTheory.Noetherian.Basic /-! # Regular sequences and weakly regular sequences diff --git a/Mathlib/RingTheory/Unramified/Field.lean b/Mathlib/RingTheory/Unramified/Field.lean index 1c82e3e8fc22d..d2ba91b1bb3d9 100644 --- a/Mathlib/RingTheory/Unramified/Field.lean +++ b/Mathlib/RingTheory/Unramified/Field.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.FieldTheory.PurelyInseparable -import Mathlib.RingTheory.Artinian +import Mathlib.RingTheory.Artinian.Ring import Mathlib.RingTheory.LocalProperties.Basic import Mathlib.Algebra.Polynomial.Taylor import Mathlib.RingTheory.Unramified.Finite From e0a3be5c67b259b7e82f70ecdec63debe3bc23f0 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Thu, 2 Jan 2025 14:27:34 +0000 Subject: [PATCH 089/189] refactor(ModelTheory): restate iAlls for a more convenient statement (#20161) Co-authored-by: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> --- Mathlib/FieldTheory/AxGrothendieck.lean | 16 +++++++------- Mathlib/ModelTheory/Semantics.lean | 29 ++++++++++--------------- Mathlib/ModelTheory/Syntax.lean | 23 ++++++++++---------- 3 files changed, 31 insertions(+), 37 deletions(-) diff --git a/Mathlib/FieldTheory/AxGrothendieck.lean b/Mathlib/FieldTheory/AxGrothendieck.lean index af34fc5d40a21..8c95945aaddd9 100644 --- a/Mathlib/FieldTheory/AxGrothendieck.lean +++ b/Mathlib/FieldTheory/AxGrothendieck.lean @@ -111,7 +111,7 @@ noncomputable def genericPolyMapSurjOnOfInjOn [Finite ι] let f2 : Language.ring.Formula ((Σ i : ι, mons i) ⊕ (Fin 2 × ι)) := iInf l2 let injOn : Language.ring.Formula (α ⊕ Σ i : ι, mons i) := - Formula.iAlls (γ := Fin 2 × ι) id + Formula.iAlls (Fin 2 × ι) (φ.relabel (Sum.map Sum.inl (fun i => (0, i))) ⟹ φ.relabel (Sum.map Sum.inl (fun i => (1, i))) ⟹ (f1.imp f2).relabel (fun x => (Equiv.sumAssoc _ _ _).symm (Sum.inr x))) @@ -122,23 +122,23 @@ noncomputable def genericPolyMapSurjOnOfInjOn [Finite ι] let f3 : Language.ring.Formula ((Σ i : ι, mons i) ⊕ (Fin 2 × ι)) := iInf l3 let surjOn : Language.ring.Formula (α ⊕ Σ i : ι, mons i) := - Formula.iAlls (γ := ι) id + Formula.iAlls ι (Formula.imp (φ.relabel (Sum.map Sum.inl id)) <| - Formula.iExs (γ := ι) + Formula.iExs ι <| + ((φ.relabel (Sum.map Sum.inl (fun i => (0, i)))) ⊓ + (f3.relabel (fun x => (Equiv.sumAssoc _ _ _).symm (Sum.inr x)))).relabel (fun (i : (α ⊕ (Σ i : ι, mons i)) ⊕ (Fin 2 × ι)) => show ((α ⊕ (Σ i : ι, mons i)) ⊕ ι) ⊕ ι from Sum.elim (Sum.inl ∘ Sum.inl) - (fun i => if i.1 = 0 then Sum.inr i.2 else (Sum.inl (Sum.inr i.2))) i) - ((φ.relabel (Sum.map Sum.inl (fun i => (0, i)))) ⊓ - (f3.relabel (fun x => (Equiv.sumAssoc _ _ _).symm (Sum.inr x))))) + (fun i => if i.1 = 0 then Sum.inr i.2 else (Sum.inl (Sum.inr i.2))) i)) let mapsTo : Language.ring.Formula (α ⊕ Σ i : ι, mons i) := - Formula.iAlls (γ := ι) id + Formula.iAlls ι (Formula.imp (φ.relabel (Sum.map Sum.inl id)) (φ.subst <| Sum.elim (fun a => .var (Sum.inl (Sum.inl a))) (fun i => (termOfFreeCommRing (genericPolyMap mons i)).relabel (fun i => (Equiv.sumAssoc _ _ _).symm (Sum.inr i))))) - Formula.iAlls (γ := α ⊕ Σ i : ι, mons i) Sum.inr (mapsTo ⟹ injOn ⟹ surjOn) + Formula.iAlls (α ⊕ Σ i : ι, mons i) ((mapsTo.imp <| injOn.imp <| surjOn).relabel Sum.inr) theorem realize_genericPolyMapSurjOnOfInjOn [Fintype ι] (φ : ring.Formula (α ⊕ ι)) (mons : ι → Finset (ι →₀ ℕ)) : diff --git a/Mathlib/ModelTheory/Semantics.lean b/Mathlib/ModelTheory/Semantics.lean index 558f312224a96..01e220cf30c65 100644 --- a/Mathlib/ModelTheory/Semantics.lean +++ b/Mathlib/ModelTheory/Semantics.lean @@ -769,10 +769,9 @@ theorem realize_exs {φ : L.BoundedFormula α n} {v : α → M} : @[simp] theorem _root_.FirstOrder.Language.Formula.realize_iAlls - [Finite γ] {f : α → β ⊕ γ} - {φ : L.Formula α} {v : β → M} : (φ.iAlls f).Realize v ↔ - ∀ (i : γ → M), φ.Realize (fun a => Sum.elim v i (f a)) := by - let e := Classical.choice (Classical.choose_spec (Finite.exists_equiv_fin γ)) + [Finite β] {φ : L.Formula (α ⊕ β)} {v : α → M} : (φ.iAlls β).Realize v ↔ + ∀ (i : β → M), φ.Realize (fun a => Sum.elim v i a) := by + let e := Classical.choice (Classical.choose_spec (Finite.exists_equiv_fin β)) rw [Formula.iAlls] simp only [Nat.add_zero, realize_alls, realize_relabel, Function.comp_def, castAdd_zero, finCongr_refl, OrderIso.refl_apply, Sum.elim_map, id_eq] @@ -787,23 +786,20 @@ theorem _root_.FirstOrder.Language.Formula.realize_iAlls exact i.elim0 @[simp] -theorem realize_iAlls [Finite γ] {f : α → β ⊕ γ} - {φ : L.Formula α} {v : β → M} {v' : Fin 0 → M} : - BoundedFormula.Realize (φ.iAlls f) v v' ↔ - ∀ (i : γ → M), φ.Realize (fun a => Sum.elim v i (f a)) := by +theorem realize_iAlls [Finite β] {φ : L.Formula (α ⊕ β)} {v : α → M} {v' : Fin 0 → M} : + BoundedFormula.Realize (φ.iAlls β) v v' ↔ + ∀ (i : β → M), φ.Realize (fun a => Sum.elim v i a) := by rw [← Formula.realize_iAlls, iff_iff_eq]; congr; simp [eq_iff_true_of_subsingleton] @[simp] theorem _root_.FirstOrder.Language.Formula.realize_iExs - [Finite γ] {f : α → β ⊕ γ} - {φ : L.Formula α} {v : β → M} : (φ.iExs f).Realize v ↔ - ∃ (i : γ → M), φ.Realize (fun a => Sum.elim v i (f a)) := by + [Finite γ] {φ : L.Formula (α ⊕ γ)} {v : α → M} : (φ.iExs γ).Realize v ↔ + ∃ (i : γ → M), φ.Realize (Sum.elim v i) := by let e := Classical.choice (Classical.choose_spec (Finite.exists_equiv_fin γ)) rw [Formula.iExs] simp only [Nat.add_zero, realize_exs, realize_relabel, Function.comp_def, castAdd_zero, finCongr_refl, OrderIso.refl_apply, Sum.elim_map, id_eq] - rw [← not_iff_not, not_exists, not_exists] - refine Equiv.forall_congr ?_ ?_ + refine Equiv.exists_congr ?_ ?_ · exact ⟨fun v => v ∘ e, fun v => v ∘ e.symm, fun _ => by simp [Function.comp_def], fun _ => by simp [Function.comp_def]⟩ @@ -814,10 +810,9 @@ theorem _root_.FirstOrder.Language.Formula.realize_iExs exact i.elim0 @[simp] -theorem realize_iExs [Finite γ] {f : α → β ⊕ γ} - {φ : L.Formula α} {v : β → M} {v' : Fin 0 → M} : - BoundedFormula.Realize (φ.iExs f) v v' ↔ - ∃ (i : γ → M), φ.Realize (fun a => Sum.elim v i (f a)) := by +theorem realize_iExs [Finite γ] {φ : L.Formula (α ⊕ γ)} {v : α → M} {v' : Fin 0 → M} : + BoundedFormula.Realize (φ.iExs γ) v v' ↔ + ∃ (i : γ → M), φ.Realize (Sum.elim v i) := by rw [← Formula.realize_iExs, iff_iff_eq]; congr; simp [eq_iff_true_of_subsingleton] @[simp] diff --git a/Mathlib/ModelTheory/Syntax.lean b/Mathlib/ModelTheory/Syntax.lean index 7baeb03e64e5d..cff9e39352aa6 100644 --- a/Mathlib/ModelTheory/Syntax.lean +++ b/Mathlib/ModelTheory/Syntax.lean @@ -744,21 +744,20 @@ protected nonrec abbrev not (φ : L.Formula α) : L.Formula α := protected abbrev imp : L.Formula α → L.Formula α → L.Formula α := BoundedFormula.imp -/-- Given a map `f : α → β ⊕ γ`, `iAlls f φ` transforms a `L.Formula α` -into a `L.Formula β` by renaming variables with the map `f` and then universally + +variable (β) in +/-- `iAlls f φ` transforms a `L.Formula (α ⊕ β)` into a `L.Formula β` by universally quantifying over all variables `Sum.inr _`. -/ -noncomputable def iAlls [Finite γ] (f : α → β ⊕ γ) - (φ : L.Formula α) : L.Formula β := - let e := Classical.choice (Classical.choose_spec (Finite.exists_equiv_fin γ)) - (BoundedFormula.relabel (fun a => Sum.map id e (f a)) φ).alls +noncomputable def iAlls [Finite β] (φ : L.Formula (α ⊕ β)) : L.Formula α := + let e := Classical.choice (Classical.choose_spec (Finite.exists_equiv_fin β)) + (BoundedFormula.relabel (fun a => Sum.map id e a) φ).alls -/-- Given a map `f : α → β ⊕ γ`, `iExs f φ` transforms a `L.Formula α` -into a `L.Formula β` by renaming variables with the map `f` and then existentially +variable (β) in +/-- `iExs f φ` transforms a `L.Formula (α ⊕ β)` into a `L.Formula β` by existentially quantifying over all variables `Sum.inr _`. -/ -noncomputable def iExs [Finite γ] (f : α → β ⊕ γ) - (φ : L.Formula α) : L.Formula β := - let e := Classical.choice (Classical.choose_spec (Finite.exists_equiv_fin γ)) - (BoundedFormula.relabel (fun a => Sum.map id e (f a)) φ).exs +noncomputable def iExs [Finite β] (φ : L.Formula (α ⊕ β)) : L.Formula α := + let e := Classical.choice (Classical.choose_spec (Finite.exists_equiv_fin β)) + (BoundedFormula.relabel (fun a => Sum.map id e a) φ).exs /-- The biimplication between formulas, as a formula. -/ protected nonrec abbrev iff (φ ψ : L.Formula α) : L.Formula α := From b65f33f1b895ed20b7ac39ce507cf3ba5faeb6d5 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 2 Jan 2025 17:45:51 +0000 Subject: [PATCH 090/189] =?UTF-8?q?feat:=20add=20`f=20=3Do[=F0=9D=95=9C;l]?= =?UTF-8?q?=20g`=20notation=20for=20`IsLittleOTVS`=20(#20305)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This makes it easier to copy over API using `f =o[l] g` from `IsLittleO`. This also changes the argument order to be consistent with `IsLittleO`. --- Mathlib/Analysis/Asymptotics/TVS.lean | 23 ++++++++++++--------- Mathlib/Analysis/Calculus/FDeriv/Basic.lean | 7 ++++--- 2 files changed, 17 insertions(+), 13 deletions(-) diff --git a/Mathlib/Analysis/Asymptotics/TVS.lean b/Mathlib/Analysis/Asymptotics/TVS.lean index 3992e19004eed..382c9bdb85964 100644 --- a/Mathlib/Analysis/Asymptotics/TVS.lean +++ b/Mathlib/Analysis/Asymptotics/TVS.lean @@ -60,7 +60,7 @@ open scoped Topology Pointwise ENNReal NNReal namespace Asymptotics -/-- `IsLittleOTVS 𝕜 f g l` is a generalization of `f =o[l] g` (`IsLittleO f g l`) +/-- `f =o[𝕜;l] g` (`IsLittleOTVS 𝕜 l f g`) is a generalization of `f =o[l] g` (`IsLittleO l f g`) that works in topological `𝕜`-vector spaces. Given two functions `f` and `g` taking values in topological vector spaces @@ -74,10 +74,13 @@ We use an `ENNReal`-valued function `egauge` for the gauge, so we unfold the definition of little o instead of reusing it. -/ def IsLittleOTVS (𝕜 : Type*) {α E F : Type*} [NNNorm 𝕜] [TopologicalSpace E] [TopologicalSpace F] [Zero E] [Zero F] [SMul 𝕜 E] [SMul 𝕜 F] - (f : α → E) (g : α → F) (l : Filter α) : Prop := + (l : Filter α) (f : α → E) (g : α → F) : Prop := ∀ U ∈ 𝓝 (0 : E), ∃ V ∈ 𝓝 (0 : F), ∀ ε ≠ (0 : ℝ≥0), ∀ᶠ x in l, egauge 𝕜 U (f x) ≤ ε * egauge 𝕜 V (g x) +@[inherit_doc] +notation:100 f " =o[" 𝕜 ";" l "] " g:100 => IsLittleOTVS 𝕜 l f g + variable {α β 𝕜 E F : Type*} section TopologicalSpace @@ -89,7 +92,7 @@ variable [NontriviallyNormedField 𝕜] theorem _root_.Filter.HasBasis.isLittleOTVS_iff {ιE ιF : Sort*} {pE : ιE → Prop} {pF : ιF → Prop} {sE : ιE → Set E} {sF : ιF → Set F} (hE : HasBasis (𝓝 (0 : E)) pE sE) (hF : HasBasis (𝓝 (0 : F)) pF sF) {f : α → E} {g : α → F} {l : Filter α} : - IsLittleOTVS 𝕜 f g l ↔ ∀ i, pE i → ∃ j, pF j ∧ ∀ ε ≠ (0 : ℝ≥0), + f =o[𝕜;l] g ↔ ∀ i, pE i → ∃ j, pF j ∧ ∀ ε ≠ (0 : ℝ≥0), ∀ᶠ x in l, egauge 𝕜 (sE i) (f x) ≤ ε * egauge 𝕜 (sF j) (g x) := by refine (hE.forall_iff ?_).trans <| forall₂_congr fun _ _ ↦ hF.exists_iff ?_ · rintro s t hsub ⟨V, hV₀, hV⟩ @@ -99,12 +102,12 @@ theorem _root_.Filter.HasBasis.isLittleOTVS_iff {ιE ιF : Sort*} {pE : ιE → @[simp] theorem isLittleOTVS_map {f : α → E} {g : α → F} {k : β → α} {l : Filter β} : - IsLittleOTVS 𝕜 f g (map k l) ↔ IsLittleOTVS 𝕜 (f ∘ k) (g ∘ k) l := by + f =o[𝕜; map k l] g ↔ (f ∘ k) =o[𝕜;l] (g ∘ k) := by simp [IsLittleOTVS] protected lemma IsLittleOTVS.smul_left {f : α → E} {g : α → F} {l : Filter α} - (h : IsLittleOTVS 𝕜 f g l) (c : α → 𝕜) : - IsLittleOTVS 𝕜 (fun x ↦ c x • f x) (fun x ↦ c x • g x) l := by + (h : f =o[𝕜;l] g) (c : α → 𝕜) : + (fun x ↦ c x • f x) =o[𝕜;l] (fun x ↦ c x • g x) := by unfold IsLittleOTVS at * peel h with U hU V hV ε hε x hx rw [egauge_smul_right, egauge_smul_right, mul_left_comm] @@ -112,7 +115,7 @@ protected lemma IsLittleOTVS.smul_left {f : α → E} {g : α → F} {l : Filter all_goals exact fun _ ↦ Filter.nonempty_of_mem ‹_› lemma isLittleOTVS_one [ContinuousSMul 𝕜 E] {f : α → E} {l : Filter α} : - IsLittleOTVS 𝕜 f (1 : α → 𝕜) l ↔ Tendsto f l (𝓝 0) := by + f =o[𝕜;l] (1 : α → 𝕜) ↔ Tendsto f l (𝓝 0) := by constructor · intro hf rw [(basis_sets _).isLittleOTVS_iff nhds_basis_ball] at hf @@ -147,7 +150,7 @@ lemma isLittleOTVS_one [ContinuousSMul 𝕜 E] {f : α → E} {l : Filter α} : simpa using le_egauge_ball_one 𝕜 (1 : 𝕜) lemma IsLittleOTVS.tendsto_inv_smul [ContinuousSMul 𝕜 E] {f : α → 𝕜} {g : α → E} {l : Filter α} - (h : IsLittleOTVS 𝕜 g f l) : Tendsto (fun x ↦ (f x)⁻¹ • g x) l (𝓝 0) := by + (h : g =o[𝕜;l] f) : Tendsto (fun x ↦ (f x)⁻¹ • g x) l (𝓝 0) := by rw [(basis_sets _).isLittleOTVS_iff nhds_basis_ball] at h rw [(nhds_basis_balanced 𝕜 E).tendsto_right_iff] rintro U ⟨hU, hUB⟩ @@ -173,7 +176,7 @@ lemma IsLittleOTVS.tendsto_inv_smul [ContinuousSMul 𝕜 E] {f : α → 𝕜} {g lemma isLittleOTVS_iff_tendsto_inv_smul [ContinuousSMul 𝕜 E] {f : α → 𝕜} {g : α → E} {l : Filter α} (h₀ : ∀ᶠ x in l, f x = 0 → g x = 0) : - IsLittleOTVS 𝕜 g f l ↔ Tendsto (fun x ↦ (f x)⁻¹ • g x) l (𝓝 0) := by + g =o[𝕜;l] f ↔ Tendsto (fun x ↦ (f x)⁻¹ • g x) l (𝓝 0) := by refine ⟨IsLittleOTVS.tendsto_inv_smul, fun h U hU ↦ ?_⟩ refine ⟨ball 0 1, ball_mem_nhds _ one_pos, fun ε hε ↦ ?_⟩ rcases NormedField.exists_norm_lt 𝕜 hε.bot_lt with ⟨c, hc₀, hcε : ‖c‖₊ < ε⟩ @@ -196,7 +199,7 @@ variable [NontriviallyNormedField 𝕜] variable [SeminormedAddCommGroup E] [SeminormedAddCommGroup F] [NormedSpace 𝕜 E] [NormedSpace 𝕜 F] lemma isLittleOTVS_iff_isLittleO {f : α → E} {g : α → F} {l : Filter α} : - IsLittleOTVS 𝕜 f g l ↔ f =o[l] g := by + f =o[𝕜;l] g ↔ f =o[l] g := by rcases NormedField.exists_one_lt_norm 𝕜 with ⟨c, hc : 1 < ‖c‖₊⟩ have hc₀ : 0 < ‖c‖₊ := one_pos.trans hc simp only [isLittleO_iff, nhds_basis_ball.isLittleOTVS_iff nhds_basis_ball] diff --git a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean index be3bc501cfa89..c22d0f09c8255 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean @@ -137,7 +137,7 @@ the notion of Fréchet derivative along the set `s`. -/ @[mk_iff hasFDerivAtFilter_iff_isLittleOTVS] structure HasFDerivAtFilter (f : E → F) (f' : E →L[𝕜] F) (x : E) (L : Filter E) : Prop where of_isLittleOTVS :: - isLittleOTVS : IsLittleOTVS 𝕜 (fun x' => f x' - f x - f' (x' - x)) (fun x' => x' - x) L + isLittleOTVS : (fun x' => f x' - f x - f' (x' - x)) =o[𝕜;L] (fun x' => x' - x) /-- A function `f` has the continuous linear map `f'` as derivative at `x` within a set `s` if `f x' = f x + f' (x' - x) + o (x' - x)` when `x'` tends to `x` inside `s`. -/ @@ -158,8 +158,9 @@ differentiable but this definition works, e.g., for vector spaces over `p`-adic @[fun_prop, mk_iff hasStrictFDerivAt_iff_isLittleOTVS] structure HasStrictFDerivAt (f : E → F) (f' : E →L[𝕜] F) (x : E) where of_isLittleOTVS :: - isLittleOTVS : IsLittleOTVS 𝕜 - (fun p : E × E => f p.1 - f p.2 - f' (p.1 - p.2)) (fun p : E × E => p.1 - p.2) (𝓝 (x, x)) + isLittleOTVS : + (fun p : E × E => f p.1 - f p.2 - f' (p.1 - p.2)) + =o[𝕜;𝓝 (x, x)] (fun p : E × E => p.1 - p.2) variable (𝕜) From f8c0388c802e232fe439c8c6739787d3e7b1b286 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Thu, 2 Jan 2025 18:44:09 +0000 Subject: [PATCH 091/189] chore(Analysis/SpecialFunctions/Pow/Complex): fix lemma name (#20406) --- Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean | 2 +- Mathlib/NumberTheory/LSeries/Basic.lean | 2 +- Mathlib/NumberTheory/LSeries/Injectivity.lean | 2 +- 3 files changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean index 2d2f3a84e0de5..0407d95d6e2cc 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean @@ -234,7 +234,7 @@ theorem conj_cpow (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : conj x ^ n = conj (x theorem cpow_conj (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : x ^ conj n = conj (conj x ^ n) := by rw [conj_cpow _ _ hx, conj_conj] -lemma cpow_natCast_add_one_ne_zero (n : ℕ) (z : ℂ) : (n + 1 : ℂ) ^ z ≠ 0 := +lemma natCast_add_one_cpow_ne_zero (n : ℕ) (z : ℂ) : (n + 1 : ℂ) ^ z ≠ 0 := mt (cpow_eq_zero_iff ..).mp fun H ↦ by norm_cast at H; exact H.1 end Complex diff --git a/Mathlib/NumberTheory/LSeries/Basic.lean b/Mathlib/NumberTheory/LSeries/Basic.lean index b6186de00d049..3faf700efb403 100644 --- a/Mathlib/NumberTheory/LSeries/Basic.lean +++ b/Mathlib/NumberTheory/LSeries/Basic.lean @@ -93,7 +93,7 @@ lemma term_congr {f g : ℕ → ℂ} (h : ∀ {n}, n ≠ 0 → f n = g n) (s : lemma pow_mul_term_eq (f : ℕ → ℂ) (s : ℂ) (n : ℕ) : (n + 1) ^ s * term f s (n + 1) = f (n + 1) := by - simp [term, cpow_natCast_add_one_ne_zero n _, mul_comm (f _), mul_div_assoc'] + simp [term, natCast_add_one_cpow_ne_zero n _, mul_comm (f _), mul_div_assoc'] lemma norm_term_eq (f : ℕ → ℂ) (s : ℂ) (n : ℕ) : ‖term f s n‖ = if n = 0 then 0 else ‖f n‖ / n ^ s.re := by diff --git a/Mathlib/NumberTheory/LSeries/Injectivity.lean b/Mathlib/NumberTheory/LSeries/Injectivity.lean index aead49d62c1c3..8e514733bfa54 100644 --- a/Mathlib/NumberTheory/LSeries/Injectivity.lean +++ b/Mathlib/NumberTheory/LSeries/Injectivity.lean @@ -64,7 +64,7 @@ lemma LSeries.tendsto_cpow_mul_atTop {f : ℕ → ℂ} {n : ℕ} (h : ∀ m ≤ · exact (hm rfl).elim · simp [H, term, (n.zero_lt_succ.trans H).ne', F, cpow_mul_div_cpow_eq_div_div_cpow] have hs {x : ℝ} (hx : x ≥ y) : Summable fun m ↦ (n + 1) ^ (x : ℂ) * term f x m := by - refine (summable_mul_left_iff <| cpow_natCast_add_one_ne_zero n _).mpr <| + refine (summable_mul_left_iff <| natCast_add_one_cpow_ne_zero n _).mpr <| LSeriesSummable_of_abscissaOfAbsConv_lt_re ?_ simpa only [ofReal_re] using hay.trans_le <| EReal.coe_le_coe_iff.mpr hx -- we can write `(n+1)^x * LSeries f x` as `f (n+1)` plus the series over `F x` From f64e251f550c694cf2f1a076fad9ae7e5634a3e1 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Thu, 2 Jan 2025 19:27:49 +0000 Subject: [PATCH 092/189] chore(LinearAlgebra/Span/Basic): add toSpanNonzeroSingleton_symm_apply_smul (#20399) --- Mathlib/LinearAlgebra/Span/Basic.lean | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/Mathlib/LinearAlgebra/Span/Basic.lean b/Mathlib/LinearAlgebra/Span/Basic.lean index 997c8ed3189d4..a6c4efb340ac8 100644 --- a/Mathlib/LinearAlgebra/Span/Basic.lean +++ b/Mathlib/LinearAlgebra/Span/Basic.lean @@ -618,10 +618,15 @@ def toSpanNonzeroSingleton : R ≃ₗ[R] R ∙ x := (LinearEquiv.ofEq (range <| toSpanSingleton R M x) (R ∙ x) (span_singleton_eq_range R M x).symm) @[simp] theorem toSpanNonzeroSingleton_apply (t : R) : - LinearEquiv.toSpanNonzeroSingleton R M x h t = + toSpanNonzeroSingleton R M x h t = (⟨t • x, Submodule.smul_mem _ _ (Submodule.mem_span_singleton_self x)⟩ : R ∙ x) := by rfl +@[simp] +lemma toSpanNonzeroSingleton_symm_apply_smul (m : R ∙ x) : + (toSpanNonzeroSingleton R M x h).symm m • x = m := + congrArg Subtype.val <| apply_symm_apply (toSpanNonzeroSingleton R M x h) m + theorem toSpanNonzeroSingleton_one : LinearEquiv.toSpanNonzeroSingleton R M x h 1 = (⟨x, Submodule.mem_span_singleton_self x⟩ : R ∙ x) := by simp From 223f9c7d067e869ccc46810df893fbef60501818 Mon Sep 17 00:00:00 2001 From: Mitchell Lee Date: Thu, 2 Jan 2025 19:57:23 +0000 Subject: [PATCH 093/189] feat: the product of two sums in a nonarchimedean ring (#12618) We prove that in a nonarchimedean ring, $$\sum_{(i, j) \in I \times J} a_i b_j = \left(\sum_{i \in I} a_i \right)\left(\sum_{j \in J} b_j\right),$$ provided that both sums on the right-hand side converge unconditionally. Also tidy some redundant `()`s from previous PRs. Co-authored-by: ADedecker --- .../Algebra/InfiniteSum/Nonarchimedean.lean | 50 +++++++++++++++++-- Mathlib/Topology/Algebra/Monoid.lean | 28 +++++------ 2 files changed, 60 insertions(+), 18 deletions(-) diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Nonarchimedean.lean b/Mathlib/Topology/Algebra/InfiniteSum/Nonarchimedean.lean index 30d1d7e493bc0..afef581ea8ffb 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Nonarchimedean.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Nonarchimedean.lean @@ -4,15 +4,22 @@ 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 +import Mathlib.Topology.Algebra.InfiniteSum.GroupCompletion +import Mathlib.Topology.Algebra.InfiniteSum.Ring +import Mathlib.Topology.Algebra.Nonarchimedean.Completion /-! # Infinite sums and products in nonarchimedean abelian groups Let `G` be a complete nonarchimedean abelian group and let `f : α → G` be a function. We prove that -`f` is unconditionally summable if and only if `f a` tends to zero on the cofinite filter on `α`. -We also prove the analogous result in the multiplicative setting. +`f` is unconditionally summable if and only if `f a` tends to zero on the cofinite filter on `α` +(`NonarchimedeanAddGroup.summable_iff_tendsto_cofinite_zero`). We also prove the analogous result in +the multiplicative setting (`NonarchimedeanGroup.multipliable_iff_tendsto_cofinite_one`). + +We also prove that multiplication distributes over arbitrarily indexed sums in a nonarchimedean +ring. That is, let `R` be a nonarchimedean ring, let `f : α → R` be a function that sums to `a : R`, +and let `g : β → R` be a function that sums to `b : R`. Then `fun (i : α × β) ↦ (f i.1) * (g i.2)` +sums to `a * b` (`HasSum.mul_of_nonarchimedean`). -/ @@ -67,3 +74,38 @@ theorem multipliable_iff_tendsto_cofinite_one [CompleteSpace G] (f : α → G) : ⟨Multipliable.tendsto_cofinite_one, multipliable_of_tendsto_cofinite_one⟩ end NonarchimedeanGroup + +section NonarchimedeanRing + +variable {α β R : Type*} +variable [Ring R] [UniformSpace R] [UniformAddGroup R] [NonarchimedeanRing R] + +/- Let `R` be a complete nonarchimedean ring. If functions `f : α → R` and `g : β → R` are summable, +then so is `fun i : α × β ↦ f i.1 * g i.2`. We will prove later that the assumption that `R` +is complete is not necessary. -/ +private theorem Summable.mul_of_complete_nonarchimedean [CompleteSpace R] {f : α → R} {g : β → R} + (hf : Summable f) (hg : Summable g) : Summable (fun i : α × β ↦ f i.1 * g i.2) := by + rw [NonarchimedeanAddGroup.summable_iff_tendsto_cofinite_zero] at * + exact tendsto_mul_cofinite_nhds_zero hf hg + +/-- Let `R` be a nonarchimedean ring, let `f : α → R` be a function that sums to `a : R`, +and let `g : β → R` be a function that sums to `b : R`. Then `fun i : α × β ↦ f i.1 * g i.2` +sums to `a * b`. -/ +theorem HasSum.mul_of_nonarchimedean {f : α → R} {g : β → R} {a b : R} (hf : HasSum f a) + (hg : HasSum g b) : HasSum (fun i : α × β ↦ f i.1 * g i.2) (a * b) := by + rw [← hasSum_iff_hasSum_compl] at * + simp only [Function.comp_def, UniformSpace.Completion.toCompl_apply, + UniformSpace.Completion.coe_mul] + exact (hf.mul hg) (hf.summable.mul_of_complete_nonarchimedean hg.summable :) + +/-- Let `R` be a nonarchimedean ring. If functions `f : α → R` and `g : β → R` are summable, then +so is `fun i : α × β ↦ f i.1 * g i.2`. -/ +theorem Summable.mul_of_nonarchimedean {f : α → R} {g : β → R} (hf : Summable f) + (hg : Summable g) : Summable (fun i : α × β ↦ f i.1 * g i.2) := + (hf.hasSum.mul_of_nonarchimedean hg.hasSum).summable + +theorem tsum_mul_tsum_of_nonarchimedean [T0Space R] {f : α → R} {g : β → R} (hf : Summable f) + (hg : Summable g) : (∑' i, f i) * (∑' i, g i) = ∑' i : α × β, f i.1 * g i.2 := + (hf.hasSum.mul_of_nonarchimedean hg.hasSum).tsum_eq.symm + +end NonarchimedeanRing diff --git a/Mathlib/Topology/Algebra/Monoid.lean b/Mathlib/Topology/Algebra/Monoid.lean index ea4a62c0da6b2..ccd647dbad6f8 100644 --- a/Mathlib/Topology/Algebra/Monoid.lean +++ b/Mathlib/Topology/Algebra/Monoid.lean @@ -430,9 +430,9 @@ Let `l` be a filter on `M` which is disjoint from the cocompact filter. Then, th `M × M → M` tends to zero on the filter product `𝓝 0 ×ˢ l`. -/ theorem tendsto_mul_nhds_zero_prod_of_disjoint_cocompact {l : Filter M} (hl : Disjoint l (cocompact M)) : - Tendsto (fun (x : M × M) ↦ x.1 * x.2) (𝓝 0 ×ˢ l) (𝓝 0) := calc - map (fun (x : M × M) ↦ x.1 * x.2) (𝓝 0 ×ˢ l) - _ ≤ map (fun (x : M × M) ↦ x.1 * x.2) (𝓝ˢ ({0} ×ˢ Set.univ)) := + Tendsto (fun x : M × M ↦ x.1 * x.2) (𝓝 0 ×ˢ l) (𝓝 0) := calc + map (fun x : M × M ↦ x.1 * x.2) (𝓝 0 ×ˢ l) + _ ≤ map (fun x : M × M ↦ x.1 * x.2) (𝓝ˢ ({0} ×ˢ Set.univ)) := map_mono <| nhds_prod_le_of_disjoint_cocompact 0 hl _ ≤ 𝓝 0 := continuous_mul.tendsto_nhdsSet_nhds fun _ ⟨hx, _⟩ ↦ mul_eq_zero_of_left hx _ @@ -441,9 +441,9 @@ Let `l` be a filter on `M` which is disjoint from the cocompact filter. Then, th `M × M → M` tends to zero on the filter product `l ×ˢ 𝓝 0`. -/ theorem tendsto_mul_prod_nhds_zero_of_disjoint_cocompact {l : Filter M} (hl : Disjoint l (cocompact M)) : - Tendsto (fun (x : M × M) ↦ x.1 * x.2) (l ×ˢ 𝓝 0) (𝓝 0) := calc - map (fun (x : M × M) ↦ x.1 * x.2) (l ×ˢ 𝓝 0) - _ ≤ map (fun (x : M × M) ↦ x.1 * x.2) (𝓝ˢ (Set.univ ×ˢ {0})) := + Tendsto (fun x : M × M ↦ x.1 * x.2) (l ×ˢ 𝓝 0) (𝓝 0) := calc + map (fun x : M × M ↦ x.1 * x.2) (l ×ˢ 𝓝 0) + _ ≤ map (fun x : M × M ↦ x.1 * x.2) (𝓝ˢ (Set.univ ×ˢ {0})) := map_mono <| prod_nhds_le_of_disjoint_cocompact 0 hl _ ≤ 𝓝 0 := continuous_mul.tendsto_nhdsSet_nhds fun _ ⟨_, hx⟩ ↦ mul_eq_zero_of_right _ hx @@ -452,7 +452,7 @@ Let `l` be a filter on `M × M` which is disjoint from the cocompact filter. The map `M × M → M` tends to zero on `(𝓝 0).coprod (𝓝 0) ⊓ l`. -/ theorem tendsto_mul_coprod_nhds_zero_inf_of_disjoint_cocompact {l : Filter (M × M)} (hl : Disjoint l (cocompact (M × M))) : - Tendsto (fun (x : M × M) ↦ x.1 * x.2) ((𝓝 0).coprod (𝓝 0) ⊓ l) (𝓝 0) := by + Tendsto (fun x : M × M ↦ x.1 * x.2) ((𝓝 0).coprod (𝓝 0) ⊓ l) (𝓝 0) := by have := calc (𝓝 0).coprod (𝓝 0) ⊓ l _ ≤ (𝓝 0).coprod (𝓝 0) ⊓ map Prod.fst l ×ˢ map Prod.snd l := @@ -470,13 +470,13 @@ Let `l` be a filter on `M × M` which is both disjoint from the cocompact filter equal to `(𝓝 0).coprod (𝓝 0)`. Then the multiplication map `M × M → M` tends to zero on `l`. -/ theorem tendsto_mul_nhds_zero_of_disjoint_cocompact {l : Filter (M × M)} (hl : Disjoint l (cocompact (M × M))) (h'l : l ≤ (𝓝 0).coprod (𝓝 0)) : - Tendsto (fun (x : M × M) ↦ x.1 * x.2) l (𝓝 0) := by + Tendsto (fun x : M × M ↦ x.1 * x.2) l (𝓝 0) := by simpa [inf_eq_right.mpr h'l] using tendsto_mul_coprod_nhds_zero_inf_of_disjoint_cocompact hl /-- Let `M` be a topological space with a continuous multiplication operation and a `0`. Let `f : α → M` and `g : α → M` be functions. If `f` tends to zero on a filter `l` and the image of `l` under `g` is disjoint from the cocompact filter on `M`, then -`fun (x : α) ↦ f x * g x` also tends to zero on `l`. -/ +`fun x : α ↦ f x * g x` also tends to zero on `l`. -/ theorem Tendsto.tendsto_mul_zero_of_disjoint_cocompact_right {f g : α → M} {l : Filter α} (hf : Tendsto f l (𝓝 0)) (hg : Disjoint (map g l) (cocompact M)) : Tendsto (fun x ↦ f x * g x) l (𝓝 0) := @@ -485,18 +485,18 @@ theorem Tendsto.tendsto_mul_zero_of_disjoint_cocompact_right {f g : α → M} {l /-- Let `M` be a topological space with a continuous multiplication operation and a `0`. Let `f : α → M` and `g : α → M` be functions. If `g` tends to zero on a filter `l` and the image of `l` under `f` is disjoint from the cocompact filter on `M`, then -`fun (x : α) ↦ f x * g x` also tends to zero on `l`. -/ +`fun x : α ↦ f x * g x` also tends to zero on `l`. -/ theorem Tendsto.tendsto_mul_zero_of_disjoint_cocompact_left {f g : α → M} {l : Filter α} (hf : Disjoint (map f l) (cocompact M)) (hg : Tendsto g l (𝓝 0)): Tendsto (fun x ↦ f x * g x) l (𝓝 0) := tendsto_mul_prod_nhds_zero_of_disjoint_cocompact hf |>.comp (tendsto_map.prod_mk hg) /-- If `f : α → M` and `g : β → M` are continuous and both tend to zero on the cocompact filter, -then `fun (i : α × β) ↦ (f i.1) * (g i.2)` also tends to zero on the cocompact filter. -/ +then `fun i : α × β ↦ f i.1 * g i.2` also tends to zero on the cocompact filter. -/ theorem tendsto_mul_cocompact_nhds_zero [TopologicalSpace α] [TopologicalSpace β] {f : α → M} {g : β → M} (f_cont : Continuous f) (g_cont : Continuous g) (hf : Tendsto f (cocompact α) (𝓝 0)) (hg : Tendsto g (cocompact β) (𝓝 0)) : - Tendsto (fun (i : α × β) ↦ (f i.1) * (g i.2)) (cocompact (α × β)) (𝓝 0) := by + Tendsto (fun i : α × β ↦ f i.1 * g i.2) (cocompact (α × β)) (𝓝 0) := by set l : Filter (M × M) := map (Prod.map f g) (cocompact (α × β)) with l_def set K : Set (M × M) := (insert 0 (range f)) ×ˢ (insert 0 (range g)) have K_compact : IsCompact K := .prod (hf.isCompact_insert_range_of_cocompact f_cont) @@ -512,10 +512,10 @@ theorem tendsto_mul_cocompact_nhds_zero [TopologicalSpace α] [TopologicalSpace exact tendsto_mul_nhds_zero_of_disjoint_cocompact l_compact l_le_coprod |>.comp tendsto_map /-- If `f : α → M` and `g : β → M` both tend to zero on the cofinite filter, then so does -`fun (i : α × β) ↦ (f i.1) * (g i.2)`. -/ +`fun i : α × β ↦ f i.1 * g i.2`. -/ theorem tendsto_mul_cofinite_nhds_zero {f : α → M} {g : β → M} (hf : Tendsto f cofinite (𝓝 0)) (hg : Tendsto g cofinite (𝓝 0)) : - Tendsto (fun (i : α × β) ↦ (f i.1) * (g i.2)) cofinite (𝓝 0) := by + Tendsto (fun i : α × β ↦ f i.1 * g i.2) cofinite (𝓝 0) := by letI : TopologicalSpace α := ⊥ haveI : DiscreteTopology α := discreteTopology_bot α letI : TopologicalSpace β := ⊥ From cc5b32e62f80c1791fcebecc76ce651ba1123a0b Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 2 Jan 2025 19:57:24 +0000 Subject: [PATCH 094/189] feat(Algebra/Colimit): the directed system of finitely generated submodules (#20264) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We show that every module is the direct limit of its finitely generated submodules. As a consequence of this and the fact that tensor products preserves colimits, we show that if `M` and `P` are arbitrary modules and `N` is a finitely generated submodule of a module `P`, then two elements of `N ⊗ M` have the same image in `P ⊗ M` if and only if they already have the same image in `N' ⊗ M` for some finitely generated submodule `N' ≥ N`. This is the theorem `Submodule.FG.exists_rTensor_fg_inclusion_eq`. --- Mathlib.lean | 2 + Mathlib/Algebra/Colimit/Finiteness.lean | 59 +++++++++++++++++++ Mathlib/Algebra/Colimit/Module.lean | 18 ++++-- Mathlib/Algebra/Colimit/Ring.lean | 6 +- Mathlib/Algebra/Colimit/TensorProduct.lean | 35 +++++++++++ .../TensorProduct/DirectLimit.lean | 18 ++++++ 6 files changed, 129 insertions(+), 9 deletions(-) create mode 100644 Mathlib/Algebra/Colimit/Finiteness.lean create mode 100644 Mathlib/Algebra/Colimit/TensorProduct.lean diff --git a/Mathlib.lean b/Mathlib.lean index 6740698922ff9..310a71920e88e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -183,8 +183,10 @@ import Mathlib.Algebra.CharZero.Infinite import Mathlib.Algebra.CharZero.Lemmas import Mathlib.Algebra.CharZero.Quotient import Mathlib.Algebra.Colimit.DirectLimit +import Mathlib.Algebra.Colimit.Finiteness import Mathlib.Algebra.Colimit.Module import Mathlib.Algebra.Colimit.Ring +import Mathlib.Algebra.Colimit.TensorProduct import Mathlib.Algebra.ContinuedFractions.Basic import Mathlib.Algebra.ContinuedFractions.Computation.ApproximationCorollaries import Mathlib.Algebra.ContinuedFractions.Computation.Approximations diff --git a/Mathlib/Algebra/Colimit/Finiteness.lean b/Mathlib/Algebra/Colimit/Finiteness.lean new file mode 100644 index 0000000000000..adb9daa65bcea --- /dev/null +++ b/Mathlib/Algebra/Colimit/Finiteness.lean @@ -0,0 +1,59 @@ +/- +Copyright (c) 2024 Junyan Xu. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Junyan Xu +-/ +import Mathlib.Algebra.Colimit.Module +import Mathlib.RingTheory.Finiteness.Basic + +/-! +# Modules as direct limits of finitely generated submodules + +We show that every module is the direct limit of its finitely generated submodules. + +## Main definitions + +* `Module.fgSystem`: the directed system of finitely generated submodules of a module. + +* `Module.fgSystem.equiv`: the isomorphism between a module and the direct limit of its +finitely generated submodules. +-/ + +namespace Module + +variable (R M : Type*) [Semiring R] [AddCommMonoid M] [Module R M] + +/-- The directed system of finitely generated submodules of a module. -/ +def fgSystem (N₁ N₂ : {N : Submodule R M // N.FG}) (le : N₁ ≤ N₂) : N₁ →ₗ[R] N₂ := + Submodule.inclusion le + +open DirectLimit + +namespace fgSystem + +instance : IsDirected {N : Submodule R M // N.FG} (· ≤ ·) where + directed N₁ N₂ := + ⟨⟨_, N₁.2.sup N₂.2⟩, Subtype.coe_le_coe.mp le_sup_left, Subtype.coe_le_coe.mp le_sup_right⟩ + +instance : DirectedSystem _ (fgSystem R M · · · ·) where + map_self _ _ := rfl + map_map _ _ _ _ _ _ := rfl + +variable [DecidableEq (Submodule R M)] + +open Submodule in +/-- Every module is the direct limit of its finitely generated submodules. -/ +noncomputable def equiv : DirectLimit _ (fgSystem R M) ≃ₗ[R] M := + .ofBijective (lift _ _ _ _ (fun _ ↦ Submodule.subtype _) fun _ _ _ _ ↦ rfl) + ⟨lift_injective _ _ fun _ ↦ Subtype.val_injective, fun x ↦ + ⟨of _ _ _ _ ⟨_, fg_span_singleton x⟩ ⟨x, subset_span <| by rfl⟩, lift_of ..⟩⟩ + +variable {R M} + +lemma equiv_comp_of (N : {N : Submodule R M // N.FG}) : + (equiv R M).toLinearMap ∘ₗ of _ _ _ _ N = N.1.subtype := by + ext; simp [equiv] + +end fgSystem + +end Module diff --git a/Mathlib/Algebra/Colimit/Module.lean b/Mathlib/Algebra/Colimit/Module.lean index 075b0220f0694..15ff35426e6c3 100644 --- a/Mathlib/Algebra/Colimit/Module.lean +++ b/Mathlib/Algebra/Colimit/Module.lean @@ -247,16 +247,22 @@ theorem linearEquiv_symm_mk {g} : (linearEquiv _ _).symm ⟦g⟧ = of _ _ G f g. end equiv -variable {G f} +variable {G f} [DirectedSystem G (f · · ·)] [IsDirected ι (· ≤ ·)] + +theorem exists_eq_of_of_eq {i x y} (h : of R ι G f i x = of R ι G f i y) : + ∃ j hij, f i j hij x = f i j hij y := by + have := Nonempty.intro i + apply_fun linearEquiv _ _ at h + simp_rw [linearEquiv_of] at h + have ⟨j, h⟩ := Quotient.exact h + exact ⟨j, h.1, h.2.2⟩ /-- A component that corresponds to zero in the direct limit is already zero in some bigger module in the directed system. -/ -theorem of.zero_exact [DirectedSystem G (f · · ·)] [IsDirected ι (· ≤ ·)] - {i x} (H : of R ι G f i x = 0) : +theorem of.zero_exact {i x} (H : of R ι G f i x = 0) : ∃ j hij, f i j hij x = (0 : G j) := by - haveI : Nonempty ι := ⟨i⟩ - apply_fun linearEquiv _ _ at H - rwa [map_zero, linearEquiv_of, DirectLimit.exists_eq_zero] at H + convert exists_eq_of_of_eq (H.trans (map_zero <| _).symm) + rw [map_zero] end DirectLimit diff --git a/Mathlib/Algebra/Colimit/Ring.lean b/Mathlib/Algebra/Colimit/Ring.lean index 0ee9c31b5230d..a1980e61705ac 100644 --- a/Mathlib/Algebra/Colimit/Ring.lean +++ b/Mathlib/Algebra/Colimit/Ring.lean @@ -158,7 +158,7 @@ theorem lift_unique (F : DirectLimit G f →+* P) (x) : F x = lift G f P (fun i ↦ F.comp <| of G f i) (fun i j hij x ↦ by simp) x := by obtain ⟨x, rfl⟩ := Ideal.Quotient.mk_surjective x exact x.induction_on (by simp) (fun _ ↦ .symm <| lift_of ..) - (by simp +contextual) (by simp+contextual) + (by simp+contextual) (by simp+contextual) lemma lift_injective [Nonempty ι] [IsDirected ι (· ≤ ·)] (injective : ∀ i, Function.Injective <| g i) : @@ -193,7 +193,7 @@ variable {G f'} bigger module in the directed system. -/ theorem of.zero_exact {i x} (hix : of G (f' · · ·) i x = 0) : ∃ (j : _) (hij : i ≤ j), f' i j hij x = 0 := by - haveI := Nonempty.intro i + have := Nonempty.intro i apply_fun ringEquiv _ _ at hix rwa [map_zero, ringEquiv_of, DirectLimit.exists_eq_zero] at hix @@ -206,7 +206,7 @@ from the components to the direct limits are injective. -/ theorem of_injective [IsDirected ι (· ≤ ·)] [DirectedSystem G fun i j h ↦ f' i j h] (hf : ∀ i j hij, Function.Injective (f' i j hij)) (i) : Function.Injective (of G (fun i j h ↦ f' i j h) i) := - haveI := Nonempty.intro i + have := Nonempty.intro i ((ringEquiv _ _).comp_injective _).mp fun _ _ eq ↦ DirectLimit.mk_injective f' hf _ (by simpa only [← ringEquiv_of]) diff --git a/Mathlib/Algebra/Colimit/TensorProduct.lean b/Mathlib/Algebra/Colimit/TensorProduct.lean new file mode 100644 index 0000000000000..17eb4f0d7d86c --- /dev/null +++ b/Mathlib/Algebra/Colimit/TensorProduct.lean @@ -0,0 +1,35 @@ +/- +Copyright (c) 2025 Junyan Xu. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Junyan Xu +-/ +import Mathlib.Algebra.Colimit.Finiteness +import Mathlib.LinearAlgebra.TensorProduct.DirectLimit + +/-! +# Tensor product with direct limit of finitely generated submodules + +we show that if `M` and `P` are arbitrary modules and `N` is a finitely generated submodule +of a module `P`, then two elements of `N ⊗ M` have the same image in `P ⊗ M` if and only if +they already have the same image in `N' ⊗ M` for some finitely generated submodule `N' ≥ N`. +This is the theorem `Submodule.FG.exists_rTensor_fg_inclusion_eq`. The key facts used are +that every module is the direct limit of its finitely generated submodules and that tensor +product preserves colimits. +-/ + +open TensorProduct + +variable {R M P : Type*} [CommSemiring R] +variable [AddCommMonoid M] [Module R M] [AddCommMonoid P] [Module R P] + +theorem Submodule.FG.exists_rTensor_fg_inclusion_eq {N : Submodule R P} (hN : N.FG) + {x y : N ⊗[R] M} (eq : N.subtype.rTensor M x = N.subtype.rTensor M y) : + ∃ N', N'.FG ∧ ∃ h : N ≤ N', (N.inclusion h).rTensor M x = (N.inclusion h).rTensor M y := by + classical + lift N to {N : Submodule R P // N.FG} using hN + apply_fun (Module.fgSystem.equiv R P).symm.toLinearMap.rTensor M at eq + apply_fun directLimitLeft _ _ at eq + simp_rw [← LinearMap.rTensor_comp_apply, ← (LinearEquiv.eq_toLinearMap_symm_comp _ _).mpr + (Module.fgSystem.equiv_comp_of N), directLimitLeft_rTensor_of] at eq + have ⟨N', le, eq⟩ := Module.DirectLimit.exists_eq_of_of_eq eq + exact ⟨_, N'.2, le, eq⟩ diff --git a/Mathlib/LinearAlgebra/TensorProduct/DirectLimit.lean b/Mathlib/LinearAlgebra/TensorProduct/DirectLimit.lean index 6af40acde4346..9c5897ca63061 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/DirectLimit.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/DirectLimit.lean @@ -89,6 +89,10 @@ noncomputable def directLimitLeft : (directLimitLeft f M).symm (of _ _ _ _ _ (g ⊗ₜ m)) = of _ _ _ f _ g ⊗ₜ m := fromDirectLimit_of_tmul f g m +lemma directLimitLeft_rTensor_of {i : ι} (x : G i ⊗[R] M) : + directLimitLeft f M (LinearMap.rTensor M (of ..) x) = of _ _ _ (f ▷ M) _ x := + x.induction_on (by simp) (by simp+contextual) (by simp+contextual) + /-- `M ⊗ (limᵢ Gᵢ)` and `limᵢ (M ⊗ Gᵢ)` are isomorphic as modules -/ @@ -106,4 +110,18 @@ noncomputable def directLimitRight : (directLimitRight f M).symm (of _ _ _ _ _ (m ⊗ₜ g)) = m ⊗ₜ of _ _ _ f _ g := by simp [directLimitRight, congr_symm_apply_of] +variable [DirectedSystem G (f · · ·)] + +instance : DirectedSystem (G · ⊗[R] M) (f ▷ M) where + map_self i x := by + convert LinearMap.rTensor_id_apply M (G i) x; ext; apply DirectedSystem.map_self' + map_map _ _ _ _ _ x := by + convert ← (LinearMap.rTensor_comp_apply M _ _ x).symm; ext; apply DirectedSystem.map_map' f + +instance : DirectedSystem (M ⊗[R] G ·) (M ◁ f) where + map_self i x := by + convert LinearMap.lTensor_id_apply M _ x; ext; apply DirectedSystem.map_self' + map_map _ _ _ h₁ h₂ x := by + convert ← (LinearMap.lTensor_comp_apply M _ _ x).symm; ext; apply DirectedSystem.map_map' f + end TensorProduct From 670870be7ff689a9e9dc6afe2fe6dec76bc5ed17 Mon Sep 17 00:00:00 2001 From: Vasily Ilin Date: Thu, 2 Jan 2025 20:19:58 +0000 Subject: [PATCH 095/189] feat(Probability/Distributions/Gaussian): the mgf of a gaussian (#19896) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Compute the mgf of a gaussian. Co-authored-by: Siyuan Ge Co-authored-by: Yaël Dillies --- .../MeasureTheory/Measure/MeasureSpace.lean | 3 ++ .../Probability/Distributions/Gaussian.lean | 32 +++++++++++++++++-- Mathlib/Probability/Moments.lean | 20 ++++++++++++ 3 files changed, 53 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index dfba5a05ca651..a389edbf53a51 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -1194,6 +1194,9 @@ theorem map_zero (f : α → β) : (0 : Measure α).map f = 0 := by theorem map_of_not_aemeasurable {f : α → β} {μ : Measure α} (hf : ¬AEMeasurable f μ) : μ.map f = 0 := by simp [map, hf] +theorem _root_.AEMeasurable.of_map_ne_zero {f : α → β} {μ : Measure α} (hf : μ.map f ≠ 0) : + AEMeasurable f μ := not_imp_comm.1 map_of_not_aemeasurable hf + theorem map_congr {f g : α → β} (h : f =ᵐ[μ] g) : Measure.map f μ = Measure.map g μ := by by_cases hf : AEMeasurable f μ · have hg : AEMeasurable g μ := hf.congr h diff --git a/Mathlib/Probability/Distributions/Gaussian.lean b/Mathlib/Probability/Distributions/Gaussian.lean index 10a5c010266d0..41df4ffe5d082 100644 --- a/Mathlib/Probability/Distributions/Gaussian.lean +++ b/Mathlib/Probability/Distributions/Gaussian.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Lorenzo Luccioli, Rémy Degenne, Alexander Bentkamp -/ import Mathlib.Analysis.SpecialFunctions.Gaussian.GaussianIntegral -import Mathlib.Probability.Notation -import Mathlib.MeasureTheory.Decomposition.Lebesgue +import Mathlib.Probability.Moments /-! # Gaussian distributions over ℝ @@ -347,6 +346,35 @@ lemma gaussianReal_mul_const {X : Ω → ℝ} (hX : Measure.map X ℙ = gaussian end Transformations +open Measurable Real + +variable {Ω : Type*} {mΩ : MeasurableSpace Ω} {p : Measure Ω} {μ : ℝ} {v : ℝ≥0} {X : Ω → ℝ} + +theorem mgf_gaussianReal (hX : p.map X = gaussianReal μ v) (t : ℝ) : + mgf X p t = exp (μ * t + v * t ^ 2 / 2) := by + by_cases hv : v = 0 + · simp only [gaussianReal, hv, ↓reduceIte] at hX + simp [mgf_dirac hX, hv] + calc + mgf X p t = (p.map X)[fun x => exp (t * x)] := by + rw [← mgf_id_map, mgf] + all_goals simp [AEMeasurable.of_map_ne_zero, hX, IsProbabilityMeasure.ne_zero] + _ = ∫ x, exp (t * x) * gaussianPDFReal μ v x := by + simp [hX, gaussianReal_of_var_ne_zero μ hv, gaussianPDF_def, ENNReal.ofReal, + integral_withDensity_eq_integral_smul (measurable_gaussianPDFReal μ v).real_toNNReal, + NNReal.smul_def, gaussianPDFReal_nonneg, mul_comm] + _ = ∫ x, exp (μ * t + v * t ^ 2 / 2) * gaussianPDFReal (μ + v * t) v x := by + simp only [gaussianPDFReal_def, mul_left_comm (exp _), mul_assoc, ← exp_add] + congr with x + field_simp only [mul_left_comm, ← exp_sub, ← exp_add] + ring_nf + _ = exp (μ * t + v * t ^ 2 / 2) := by + rw [integral_mul_left, integral_gaussianPDFReal_eq_one (μ + v * t) hv, mul_one] + +theorem cgf_gaussianReal (hX : p.map X = gaussianReal μ v) (t : ℝ) : + cgf X p t = μ * t + v * t ^ 2 / 2 := by + rw [cgf, mgf_gaussianReal hX t, log_exp] + end GaussianReal end ProbabilityTheory diff --git a/Mathlib/Probability/Moments.lean b/Mathlib/Probability/Moments.lean index 1178661e1715f..41fb7a9465a19 100644 --- a/Mathlib/Probability/Moments.lean +++ b/Mathlib/Probability/Moments.lean @@ -180,6 +180,12 @@ theorem mgf_pos [IsProbabilityMeasure μ] (h_int_X : Integrable (fun ω => exp ( 0 < mgf X μ t := mgf_pos' (IsProbabilityMeasure.ne_zero μ) h_int_X +lemma mgf_id_map (hX : AEMeasurable X μ) : mgf id (μ.map X) = mgf X μ := by + ext t + rw [mgf, integral_map hX] + · rfl + · exact (measurable_const_mul _).exp.aestronglyMeasurable + theorem mgf_neg : mgf (-X) μ t = mgf X μ (-t) := by simp_rw [mgf, Pi.neg_apply, mul_neg, neg_mul] theorem cgf_neg : cgf (-X) μ t = cgf X μ (-t) := by simp_rw [cgf, mgf_neg] @@ -187,6 +193,15 @@ theorem cgf_neg : cgf (-X) μ t = cgf X μ (-t) := by simp_rw [cgf, mgf_neg] theorem mgf_smul_left (α : ℝ) : mgf (α • X) μ t = mgf X μ (α * t) := by simp_rw [mgf, Pi.smul_apply, smul_eq_mul, mul_comm α t, mul_assoc] +theorem mgf_const_add (α : ℝ) : mgf (fun ω => α + X ω) μ t = exp (t * α) * mgf X μ t := by + rw [mgf, mgf, ← integral_mul_left] + congr with x + dsimp + rw [mul_add, exp_add] + +theorem mgf_add_const (α : ℝ) : mgf (fun ω => X ω + α) μ t = mgf X μ t * exp (t * α) := by + simp only [add_comm, mgf_const_add, mul_comm] + /-- This is a trivial application of `IndepFun.comp` but it will come up frequently. -/ theorem IndepFun.exp_mul {X Y : Ω → ℝ} (h_indep : IndepFun X Y μ) (s t : ℝ) : IndepFun (fun ω => exp (s * X ω)) (fun ω => exp (t * Y ω)) μ := by @@ -334,6 +349,11 @@ theorem measure_le_le_exp_cgf [IsFiniteMeasure μ] (ε : ℝ) (ht : t ≤ 0) rw [exp_add] exact mul_le_mul le_rfl (le_exp_log _) mgf_nonneg (exp_pos _).le +lemma mgf_dirac {x : ℝ} (hX : μ.map X = .dirac x) (t : ℝ) : mgf X μ t = exp (x * t) := by + have : IsProbabilityMeasure (μ.map X) := by rw [hX]; infer_instance + rw [← mgf_id_map (.of_map_ne_zero <| IsProbabilityMeasure.ne_zero _), mgf, hX, integral_dirac, + mul_comm, id_def] + end MomentGeneratingFunction end ProbabilityTheory From 0b74e84c50e04a158e97442a7fb34acb0200d621 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Thu, 2 Jan 2025 20:35:53 +0000 Subject: [PATCH 096/189] refactor(NumberTheory/Ostrowski): use `AbsoluteValue` instead of `MulRingNorm` (#20362) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR * replaces `MulRingNorm R` by `AbsoluteValue R ℝ` in the material developed in `Mathlib.NumberTheory.Ostrowski` * does some cleanup (naming convention, remove redundant name parts etc.) * and also some golfing... In addition, it * creates a new folder `Mathlib.Algebra.Order.AbsoluteValue` * moves `Mathlib.Algebra.Order.AbsoluteValue` to `Mathlib.Algebra.Order.AbsoluteValue.Basic` and `Mathlib.Algebra.Order.EuclideanAbsoluteValue`to `Mathlib.Algebra.Order.AbsoluteValue.Euclidean` * adds some API for `AbsoluteValue` to match `MulRingNorm` (in `Mathlib.Algebra.Order.AbsoluteValue.Basic`) * adds a new file `Mathlib.Algebra.Order.AbsoluteValue.Equivalence` with material on equivalence of real-valued absolute values. See the [discussion](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/MulRingNorm.20vs.2E.20AbsoluteValue/near/491324423) on Zulip. Co-authored-by: Michael Stoll <99838730+MichaelStollBayreuth@users.noreply.github.com> --- Mathlib.lean | 5 +- .../Basic.lean} | 67 ++- .../Order/AbsoluteValue/Equivalence.lean | 62 +++ .../Euclidean.lean} | 2 +- .../Order/BigOperators/Ring/Finset.lean | 2 +- Mathlib/Algebra/Order/CauSeq/Basic.lean | 2 +- .../Polynomial/Degree/CardPowDegree.lean | 2 +- Mathlib/Analysis/Normed/Ring/WithAbs.lean | 2 +- Mathlib/Data/Int/AbsoluteValue.lean | 2 +- .../ClassNumber/AdmissibleAbsoluteValue.lean | 2 +- Mathlib/NumberTheory/Ostrowski.lean | 469 ++++++++---------- Mathlib/Topology/Algebra/Ring/Basic.lean | 2 +- .../Topology/UniformSpace/AbsoluteValue.lean | 2 +- MathlibTest/ValuedCSP.lean | 2 +- 14 files changed, 354 insertions(+), 269 deletions(-) rename Mathlib/Algebra/Order/{AbsoluteValue.lean => AbsoluteValue/Basic.lean} (84%) create mode 100644 Mathlib/Algebra/Order/AbsoluteValue/Equivalence.lean rename Mathlib/Algebra/Order/{EuclideanAbsoluteValue.lean => AbsoluteValue/Euclidean.lean} (97%) diff --git a/Mathlib.lean b/Mathlib.lean index 310a71920e88e..af4002f52b026 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -628,7 +628,9 @@ import Mathlib.Algebra.NoZeroSMulDivisors.Pi import Mathlib.Algebra.NoZeroSMulDivisors.Prod import Mathlib.Algebra.Notation import Mathlib.Algebra.Opposites -import Mathlib.Algebra.Order.AbsoluteValue +import Mathlib.Algebra.Order.AbsoluteValue.Basic +import Mathlib.Algebra.Order.AbsoluteValue.Equivalence +import Mathlib.Algebra.Order.AbsoluteValue.Euclidean import Mathlib.Algebra.Order.AddGroupWithTop import Mathlib.Algebra.Order.AddTorsor import Mathlib.Algebra.Order.Algebra @@ -654,7 +656,6 @@ import Mathlib.Algebra.Order.CauSeq.BigOperators import Mathlib.Algebra.Order.CauSeq.Completion import Mathlib.Algebra.Order.Chebyshev import Mathlib.Algebra.Order.CompleteField -import Mathlib.Algebra.Order.EuclideanAbsoluteValue import Mathlib.Algebra.Order.Field.Basic import Mathlib.Algebra.Order.Field.Canonical.Basic import Mathlib.Algebra.Order.Field.Canonical.Defs diff --git a/Mathlib/Algebra/Order/AbsoluteValue.lean b/Mathlib/Algebra/Order/AbsoluteValue/Basic.lean similarity index 84% rename from Mathlib/Algebra/Order/AbsoluteValue.lean rename to Mathlib/Algebra/Order/AbsoluteValue/Basic.lean index 657398ad3fd64..c9440b72f7831 100644 --- a/Mathlib/Algebra/Order/AbsoluteValue.lean +++ b/Mathlib/Algebra/Order/AbsoluteValue/Basic.lean @@ -93,6 +93,12 @@ protected theorem eq_zero {x : R} : abv x = 0 ↔ x = 0 := protected theorem add_le (x y : R) : abv (x + y) ≤ abv x + abv y := abv.add_le' x y +/-- The triangle inequality for an `AbsoluteValue` applied to a list. -/ +lemma listSum_le (l : List R) : abv l.sum ≤ (l.map abv).sum := by + induction l with + | nil => simp + | cons head tail ih => exact (abv.add_le ..).trans <| add_le_add_left ih (abv head) + @[simp] protected theorem map_mul (x y : R) : abv (x * y) = abv x * abv y := abv.map_mul' x y @@ -174,6 +180,20 @@ theorem coe_toMonoidHom : ⇑abv.toMonoidHom = abv := protected theorem map_pow (a : R) (n : ℕ) : abv (a ^ n) = abv a ^ n := abv.toMonoidHom.map_pow a n +omit [Nontrivial R] in +/-- An absolute value satisfies `f (n : R) ≤ n` for every `n : ℕ`. -/ +lemma apply_nat_le_self (n : ℕ) : abv n ≤ n := by + cases subsingleton_or_nontrivial R + · simp [Subsingleton.eq_zero (n : R)] + induction n with + | zero => simp + | succ n hn => + simp only [Nat.cast_succ] + calc + abv (n + 1) ≤ abv n + abv 1 := abv.add_le .. + _ = abv n + 1 := congrArg (abv n + ·) abv.map_one + _ ≤ n + 1 := add_le_add_right hn 1 + end IsDomain end Semiring @@ -191,8 +211,8 @@ end Ring end OrderedRing section OrderedCommRing -variable [OrderedCommRing S] [Ring R] (abv : AbsoluteValue R S) -variable [NoZeroDivisors S] + +variable [OrderedCommRing S] [Ring R] (abv : AbsoluteValue R S) [NoZeroDivisors S] @[simp] protected theorem map_neg (a : R) : abv (-a) = abv a := by @@ -219,6 +239,18 @@ instance [Nontrivial R] [IsDomain S] : MulRingNormClass (AbsoluteValue R S) R S map_neg_eq_map := fun f => f.map_neg eq_zero_of_map_eq_zero := fun f _ => f.eq_zero.1 } +open Int in +lemma apply_natAbs_eq (x : ℤ) : abv (natAbs x) = abv x := by + obtain ⟨_, rfl | rfl⟩ := eq_nat_or_neg x <;> simp + +open Int in +/-- Values of an absolute value coincide on the image of `ℕ` in `R` +if and only if they coincide on the image of `ℤ` in `R`. -/ +lemma eq_on_nat_iff_eq_on_int {f g : AbsoluteValue R S} : + (∀ n : ℕ , f n = g n) ↔ ∀ n : ℤ , f n = g n := by + refine ⟨fun h z ↦ ?_, fun a n ↦ mod_cast a n⟩ + obtain ⟨n , rfl | rfl⟩ := eq_nat_or_neg z <;> simp [h n] + end OrderedCommRing section LinearOrderedRing @@ -249,6 +281,37 @@ theorem abs_abv_sub_le_abv_sub (a b : R) : abs (abv a - abv b) ≤ abv (a - b) : end LinearOrderedCommRing +section trivial + +variable {R : Type*} [Semiring R] [DecidablePred fun x : R ↦ x = 0] [NoZeroDivisors R] +variable {S : Type*} [OrderedSemiring S] [Nontrivial S] + +/-- The *trivial* absolute value takes the value `1` on all nonzero elements. -/ +protected +def trivial: AbsoluteValue R S where + toFun x := if x = 0 then 0 else 1 + map_mul' x y := by + rcases eq_or_ne x 0 with rfl | hx + · simp + rcases eq_or_ne y 0 with rfl | hy + · simp + simp [hx, hy] + nonneg' x := by rcases eq_or_ne x 0 with hx | hx <;> simp [hx] + eq_zero' x := by rcases eq_or_ne x 0 with hx | hx <;> simp [hx] + add_le' x y := by + rcases eq_or_ne x 0 with rfl | hx + · simp + rcases eq_or_ne y 0 with rfl | hy + · simp + simp only [hx, ↓reduceIte, hy, one_add_one_eq_two] + rcases eq_or_ne (x + y) 0 with hxy | hxy <;> simp [hxy, one_le_two] + +@[simp] +lemma trivial_apply {x : R} (hx : x ≠ 0) : AbsoluteValue.trivial (S := S) x = 1 := + if_neg hx + +end trivial + end AbsoluteValue -- Porting note: Removed [] in fields, see diff --git a/Mathlib/Algebra/Order/AbsoluteValue/Equivalence.lean b/Mathlib/Algebra/Order/AbsoluteValue/Equivalence.lean new file mode 100644 index 0000000000000..bd4fa328898cc --- /dev/null +++ b/Mathlib/Algebra/Order/AbsoluteValue/Equivalence.lean @@ -0,0 +1,62 @@ +/- +Copyright (c) 2025 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Stoll +-/ +import Mathlib.Analysis.SpecialFunctions.Pow.Real + +/-! +# Equivalence of real-valued absolute values + +Two absolute values `v₁, v₂ : AbsoluteValue R ℝ` are *equivalent* if there exists a +positive real number `c` such that `v₁ x ^ c = v₂ x` for all `x : R`. +-/ + +namespace AbsoluteValue + +variable {R : Type*} [Semiring R] + +/-- Two absolute values `f, g` on `R` with values in `ℝ` are *equivalent* if there exists +a positive real constant `c` such that for all `x : R`, `(f x)^c = g x`. -/ +def Equiv (f g : AbsoluteValue R ℝ) : Prop := + ∃ c : ℝ, 0 < c ∧ (f · ^ c) = g + +/-- Equivalence of absolute values is reflexive. -/ +lemma equiv_refl (f : AbsoluteValue R ℝ) : Equiv f f := + ⟨1, Real.zero_lt_one, funext fun x ↦ Real.rpow_one (f x)⟩ + +/-- Equivalence of absolute values is symmetric. -/ +lemma equiv_symm {f g : AbsoluteValue R ℝ} (hfg : Equiv f g) : Equiv g f := by + rcases hfg with ⟨c, hcpos, h⟩ + refine ⟨1 / c, one_div_pos.mpr hcpos, ?_⟩ + simp [← h, Real.rpow_rpow_inv (apply_nonneg f _) (ne_of_lt hcpos).symm] + +/-- Equivalence of absolute values is transitive. -/ +lemma equiv_trans {f g k : AbsoluteValue R ℝ} (hfg : Equiv f g) (hgk : Equiv g k) : + Equiv f k := by + rcases hfg with ⟨c, hcPos, hfg⟩ + rcases hgk with ⟨d, hdPos, hgk⟩ + refine ⟨c * d, mul_pos hcPos hdPos, ?_⟩ + simp [← hgk, ← hfg, Real.rpow_mul (apply_nonneg f _)] + +/-- An absolute value is equivalent to the trivial iff it is trivial itself. -/ +@[simp] +lemma eq_trivial_of_equiv_trivial [DecidablePred fun x : R ↦ x = 0] [NoZeroDivisors R] + (f : AbsoluteValue R ℝ) : + f.Equiv .trivial ↔ f = .trivial := by + refine ⟨fun ⟨c, hc₀, hc⟩ ↦ ext fun x ↦ ?_, fun H ↦ H ▸ equiv_refl f⟩ + apply_fun (· x) at hc + rcases eq_or_ne x 0 with rfl | hx + · simp + · simp only [ne_eq, hx, not_false_eq_true, trivial_apply] at hc ⊢ + exact (Real.rpow_left_inj (f.nonneg x) zero_le_one hc₀.ne').mp <| (Real.one_rpow c).symm ▸ hc + +instance : Setoid (AbsoluteValue R ℝ) where + r := Equiv + iseqv := { + refl := equiv_refl + symm := equiv_symm + trans := equiv_trans + } + +end AbsoluteValue diff --git a/Mathlib/Algebra/Order/EuclideanAbsoluteValue.lean b/Mathlib/Algebra/Order/AbsoluteValue/Euclidean.lean similarity index 97% rename from Mathlib/Algebra/Order/EuclideanAbsoluteValue.lean rename to Mathlib/Algebra/Order/AbsoluteValue/Euclidean.lean index d294a775f14d0..8233ab4119f6e 100644 --- a/Mathlib/Algebra/Order/EuclideanAbsoluteValue.lean +++ b/Mathlib/Algebra/Order/AbsoluteValue/Euclidean.lean @@ -3,7 +3,7 @@ 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.Algebra.Order.AbsoluteValue +import Mathlib.Algebra.Order.AbsoluteValue.Basic import Mathlib.Algebra.EuclideanDomain.Int /-! diff --git a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean index f3107d2d11cbf..607f6e70c1a6c 100644 --- a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ import Mathlib.Algebra.BigOperators.Ring -import Mathlib.Algebra.Order.AbsoluteValue +import Mathlib.Algebra.Order.AbsoluteValue.Basic import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Algebra.Order.BigOperators.Ring.Multiset import Mathlib.Tactic.Ring diff --git a/Mathlib/Algebra/Order/CauSeq/Basic.lean b/Mathlib/Algebra/Order/CauSeq/Basic.lean index 4dacb249f8e44..5ba2dc7d77f6f 100644 --- a/Mathlib/Algebra/Order/CauSeq/Basic.lean +++ b/Mathlib/Algebra/Order/CauSeq/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.Action.Pi -import Mathlib.Algebra.Order.AbsoluteValue +import Mathlib.Algebra.Order.AbsoluteValue.Basic import Mathlib.Algebra.Order.Field.Basic import Mathlib.Algebra.Order.Group.MinMax import Mathlib.Algebra.Ring.Pi diff --git a/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean b/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean index 5d5a02e7b4b87..57d52ad4ec5c3 100644 --- a/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean +++ b/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean @@ -3,7 +3,7 @@ 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.Algebra.Order.EuclideanAbsoluteValue +import Mathlib.Algebra.Order.AbsoluteValue.Euclidean import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Algebra.Polynomial.FieldDivision diff --git a/Mathlib/Analysis/Normed/Ring/WithAbs.lean b/Mathlib/Analysis/Normed/Ring/WithAbs.lean index 4cbcad5fc3e00..9a69e4045dd4d 100644 --- a/Mathlib/Analysis/Normed/Ring/WithAbs.lean +++ b/Mathlib/Analysis/Normed/Ring/WithAbs.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Salvatore Mercuri -/ import Mathlib.Algebra.Group.Basic -import Mathlib.Algebra.Order.AbsoluteValue +import Mathlib.Algebra.Order.AbsoluteValue.Basic import Mathlib.Analysis.Normed.Field.Basic import Mathlib.Analysis.Normed.Module.Completion diff --git a/Mathlib/Data/Int/AbsoluteValue.lean b/Mathlib/Data/Int/AbsoluteValue.lean index f37f06dffef6a..ac42e74e70507 100644 --- a/Mathlib/Data/Int/AbsoluteValue.lean +++ b/Mathlib/Data/Int/AbsoluteValue.lean @@ -5,7 +5,7 @@ Authors: Anne Baanen -/ import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.Module.Basic -import Mathlib.Algebra.Order.AbsoluteValue +import Mathlib.Algebra.Order.AbsoluteValue.Basic import Mathlib.Algebra.Ring.Int.Units import Mathlib.Data.Int.Cast.Lemmas diff --git a/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean b/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean index 0d34bda52b147..7e288de6b9c30 100644 --- a/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean +++ b/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean @@ -5,7 +5,7 @@ Authors: Anne Baanen -/ import Mathlib.Data.Real.Basic import Mathlib.Combinatorics.Pigeonhole -import Mathlib.Algebra.Order.EuclideanAbsoluteValue +import Mathlib.Algebra.Order.AbsoluteValue.Euclidean /-! # Admissible absolute values diff --git a/Mathlib/NumberTheory/Ostrowski.lean b/Mathlib/NumberTheory/Ostrowski.lean index 31c4f57ad0b8c..2deb7a4a4f251 100644 --- a/Mathlib/NumberTheory/Ostrowski.lean +++ b/Mathlib/NumberTheory/Ostrowski.lean @@ -6,10 +6,8 @@ María Inés de Frutos-Fernández, Sam van Gool, Silvain Rideau-Kikuchi, Amos Tu Francesco Veneziano -/ -import Mathlib.Analysis.Normed.Field.Lemmas -import Mathlib.Analysis.Normed.Ring.Seminorm +import Mathlib.Algebra.Order.AbsoluteValue.Equivalence import Mathlib.Analysis.SpecialFunctions.Log.Base -import Mathlib.Analysis.SpecialFunctions.Pow.Asymptotics import Mathlib.Analysis.SpecialFunctions.Pow.Continuity import Mathlib.NumberTheory.Padics.PadicNorm @@ -21,9 +19,9 @@ Ostrowski's Theorem for the field `ℚ`: every absolute value on `ℚ` is equiva ## Main results -- `mulRingNorm_equiv_standard_or_padic`: given an absolute value on `ℚ`, it is equivalent to the -standard Archimedean (Euclidean) absolute value or to a `p`-adic absolute value for some prime -number `p`. +- `Rat.AbsoluteValue.equiv_real_or_padic`: given an absolute value on `ℚ`, it is equivalent +to the standard Archimedean (Euclidean) absolute value `Rat.AbsoluteValue.real` or to a `p`-adic +absolute value `Rat.AbsoluteValue.padic p` for a unique prime number `p`. ## TODO @@ -37,124 +35,106 @@ Extend to arbitrary number fields. ## Tags -ring norm, ostrowski +absolute value, Ostrowski's theorem -/ -/- ## Preliminary lemmas on limits and lists -/ - - open Filter Nat Real Topology -- For any `C > 0`, the limit of `C ^ (1/k)` is 1 as `k → ∞` -private lemma tendsto_root_atTop_nhds_one {C : ℝ} (hC : 0 < C) : - Tendsto (fun k : ℕ ↦ C ^ (k : ℝ)⁻¹) atTop (𝓝 1) := by - convert_to Tendsto ((fun k ↦ C ^ k) ∘ (fun k : ℝ ↦ k⁻¹) ∘ (Nat.cast)) - atTop (𝓝 1) - exact Tendsto.comp (Continuous.tendsto' (continuous_iff_continuousAt.2 - (fun a ↦ continuousAt_const_rpow hC.ne')) 0 1 (rpow_zero C)) - <| Tendsto.comp tendsto_inv_atTop_zero tendsto_natCast_atTop_atTop +private lemma tendsto_const_rpow_inv {C : ℝ} (hC : 0 < C) : + Tendsto (fun k : ℕ ↦ C ^ (k : ℝ)⁻¹) atTop (𝓝 1) := + ((continuous_iff_continuousAt.mpr fun _ ↦ continuousAt_const_rpow hC.ne').tendsto' + 0 1 (rpow_zero C)).comp <| tendsto_inv_atTop_zero.comp tendsto_natCast_atTop_atTop --extends the lemma `tendsto_rpow_div` when the function has natural input -private lemma tendsto_nat_rpow_div : +private lemma tendsto_nat_rpow_inv : Tendsto (fun k : ℕ ↦ (k : ℝ) ^ (k : ℝ)⁻¹) atTop (𝓝 1) := by - convert_to Tendsto ((fun k : ℝ ↦ k ^ k⁻¹) ∘ Nat.cast) atTop (𝓝 1) - apply Tendsto.comp _ tendsto_natCast_atTop_atTop simp_rw [← one_div] - exact tendsto_rpow_div + exact Tendsto.comp tendsto_rpow_div tendsto_natCast_atTop_atTop -- Multiplication by a constant moves in a List.sum -private lemma list_mul_sum {R : Type*} [CommSemiring R] {T : Type*} (l : List T) (y : R) : ∀ x : R, - List.sum (List.mapIdx (fun i _ => x * y ^ i) (l)) = - x * List.sum (List.mapIdx (fun i _ => y ^ i) (l)) := by - induction l with - | nil => simp only [List.mapIdx_nil, List.sum_nil, mul_zero, forall_const] - | cons head tail ih => - intro x - simp_rw [List.mapIdx_cons, pow_zero, mul_one, List.sum_cons, mul_add, mul_one] - have (a : ℕ) : y ^ (a + 1) = y * y ^ a := by ring - simp_rw [this, ← mul_assoc, ih, ← mul_assoc] +private lemma list_mul_sum {R : Type*} [CommSemiring R] {T : Type*} (l : List T) (y : R) (x : R) : + (l.mapIdx fun i _ => x * y ^ i).sum = x * (l.mapIdx fun i _ => y ^ i).sum := by + simp_rw [← smul_eq_mul, List.smul_sum, List.mapIdx_eq_enum_map] + congr 1 + simp -- Geometric sum for lists private lemma list_geom {T : Type*} {F : Type*} [Field F] (l : List T) {y : F} (hy : y ≠ 1) : - List.sum (List.mapIdx (fun i _ => y ^ i) l) = (y ^ l.length - 1) / (y - 1) := by - induction l with - | nil => simp only [List.mapIdx_nil, List.sum_nil, List.length_nil, pow_zero, sub_self, zero_div] - | cons head tail ih => - simp only [List.mapIdx_cons, pow_zero, List.sum_cons, List.length_cons] - have (a : ℕ ) : y ^ (a + 1) = y * y ^ a := by ring - simp_rw [this, list_mul_sum, ih] - simp only [mul_div, ← same_add_div (sub_ne_zero.2 hy), mul_sub, mul_one, sub_add_sub_cancel'] - -namespace Rat.MulRingNorm -open Int + (l.mapIdx fun i _ => y ^ i).sum = (y ^ l.length - 1) / (y - 1) := by + rw [← geom_sum_eq hy l.length, List.mapIdx_eq_enum_map, Finset.sum_range, ← Fin.sum_univ_get'] + simp only [List.getElem_enum, Function.uncurry_apply_pair] + let e : Fin l.enum.length ≃ Fin l.length := finCongr List.enum_length + exact Fintype.sum_bijective e e.bijective _ _ fun _ ↦ rfl + +open AbsoluteValue -- does not work as intended after `namespace Rat.AbsoluteValue` + +namespace Rat.AbsoluteValue -variable {f g : MulRingNorm ℚ} +/-! +### Preliminary lemmas +-/ -/-- Values of a multiplicative norm of the rationals coincide on ℕ if and only if they coincide -on `ℤ`. -/ -lemma eq_on_nat_iff_eq_on_Int : (∀ n : ℕ , f n = g n) ↔ (∀ n : ℤ , f n = g n) := by - refine ⟨fun h z ↦ ?_, fun a n ↦ a n⟩ - obtain ⟨n , rfl | rfl⟩ := eq_nat_or_neg z <;> - simp only [Int.cast_neg, Int.cast_natCast, map_neg_eq_map, h n] +open Int -/-- Values of a multiplicative norm of the rationals are determined by the values on the natural +variable {f g : AbsoluteValue ℚ ℝ} + +/-- Values of an absolute value on the rationals are determined by the values on the natural numbers. -/ lemma eq_on_nat_iff_eq : (∀ n : ℕ , f n = g n) ↔ f = g := by refine ⟨fun h ↦ ?_, fun h n ↦ congrFun (congrArg DFunLike.coe h) ↑n⟩ - ext z - rw [← Rat.num_div_den z, map_div₀, map_div₀, h, eq_on_nat_iff_eq_on_Int.mp h] + ext1 z + rw [← Rat.num_div_den z, map_div₀, map_div₀, h, eq_on_nat_iff_eq_on_int.mp h] -/-- The equivalence class of a multiplicative norm on the rationals is determined by its values on +/-- The equivalence class of an absolute value on the rationals is determined by its values on the natural numbers. -/ -lemma equiv_on_nat_iff_equiv : (∃ c : ℝ, 0 < c ∧ (∀ n : ℕ , (f n) ^ c = g n)) ↔ - f.equiv g := by - refine ⟨fun ⟨c, hc, h⟩ ↦ ⟨c, ⟨hc, ?_⟩⟩, fun ⟨c, hc, h⟩ ↦ ⟨c, ⟨hc, fun n ↦ by rw [← h]⟩⟩⟩ - ext x - rw [← Rat.num_div_den x, map_div₀, map_div₀, div_rpow (by positivity) (by positivity), h x.den, - ← MulRingNorm.apply_natAbs_eq,← MulRingNorm.apply_natAbs_eq, h (natAbs x.num)] - -open Rat.MulRingNorm +lemma equiv_on_nat_iff_equiv : (∃ c : ℝ, 0 < c ∧ ∀ n : ℕ , f n ^ c = g n) ↔ f ≈ g := by + refine ⟨fun ⟨c, hc, h⟩ ↦ ⟨c, hc, ?_⟩, fun ⟨c, hc, h⟩ ↦ ⟨c, hc, (congrFun h ·)⟩⟩ + ext1 x + rw [← Rat.num_div_den x, map_div₀, map_div₀, div_rpow (by positivity) (by positivity), h x.den, + ← apply_natAbs_eq,← apply_natAbs_eq, h (natAbs x.num)] section Non_archimedean -/-! ## Non-archimedean case +/-! +### The non-archimedean case -Every bounded absolute value is equivalent to a `p`-adic absolute value +Every bounded absolute value on `ℚ` is equivalent to a `p`-adic absolute value. -/ -/-- The mulRingNorm corresponding to the p-adic norm on `ℚ`. -/ -def mulRingNorm_padic (p : ℕ) [Fact p.Prime] : MulRingNorm ℚ := -{ toFun := fun x : ℚ ↦ (padicNorm p x : ℝ), - map_zero' := by simp only [padicNorm.zero, Rat.cast_zero] - add_le' := by simp only; norm_cast; exact fun r s ↦ padicNorm.triangle_ineq r s - neg' := by simp only [forall_const, padicNorm.neg] - eq_zero_of_map_eq_zero' := by - simp only [Rat.cast_eq_zero] - apply padicNorm.zero_of_padicNorm_eq_zero - map_one' := by simp only [ne_eq, one_ne_zero, not_false_eq_true, padicNorm.eq_zpow_of_nonzero, - padicValRat.one, neg_zero, zpow_zero, Rat.cast_one] +/-- The real-valued `AbsoluteValue` corresponding to the p-adic norm on `ℚ`. -/ +def padic (p : ℕ) [Fact p.Prime] : AbsoluteValue ℚ ℝ where + toFun x := (padicNorm p x : ℝ) map_mul' := by simp only [padicNorm.mul, Rat.cast_mul, forall_const] -} + nonneg' x := cast_nonneg.mpr <| padicNorm.nonneg x + eq_zero' x := + ⟨fun H ↦ padicNorm.zero_of_padicNorm_eq_zero <| cast_eq_zero.mp H, + fun H ↦ cast_eq_zero.mpr <| H ▸ padicNorm.zero (p := p)⟩ + add_le' x y := by simp only; exact_mod_cast padicNorm.triangle_ineq x y + +@[simp] lemma padic_eq_padicNorm (p : ℕ) [Fact p.Prime] (r : ℚ) : + padic p r = padicNorm p r := rfl -@[simp] lemma mulRingNorm_eq_padic_norm (p : ℕ) [Fact p.Prime] (r : ℚ) : - mulRingNorm_padic p r = padicNorm p r := rfl +lemma padic_le_one (p : ℕ) [Fact p.Prime] (n : ℤ) : padic p n ≤ 1 := by + simp only [padic_eq_padicNorm] + exact_mod_cast padicNorm.of_int n -- ## Step 1: define `p = minimal n s. t. 0 < f n < 1` -variable (hf_nontriv : f ≠ 1) (bdd : ∀ n : ℕ, f n ≤ 1) +variable (hf_nontriv : f ≠ .trivial) (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 +lemma exists_minimal_nat_zero_lt_and_lt_one : + ∃ p : ℕ, (0 < f p ∧ f p < 1) ∧ ∀ m : ℕ, 0 < f m ∧ f m < 1 → p ≤ m := by -- There is a positive integer with absolute value different from one. obtain ⟨n, hn1, hn2⟩ : ∃ n : ℕ, n ≠ 0 ∧ f n ≠ 1 := by contrapose! hf_nontriv - rw [← eq_on_nat_iff_eq] + rw [AbsoluteValue.trivial, ← eq_on_nat_iff_eq] intro n rcases eq_or_ne n 0 with rfl | hn0 - · simp only [Nat.cast_zero, map_zero] - · simp only [MulRingNorm.apply_one, Nat.cast_eq_zero, hn0, ↓reduceIte, hf_nontriv n hn0] + · simp + · simp [hn0, hf_nontriv n hn0] set P := {m : ℕ | 0 < f ↑m ∧ f ↑m < 1} -- p is going to be the minimum of this set. have hP : P.Nonempty := ⟨n, map_pos_of_ne_zero f (Nat.cast_ne_zero.mpr hn1), lt_of_le_of_ne (bdd n) hn2⟩ @@ -166,7 +146,7 @@ variable {p : ℕ} (hp0 : 0 < f p) (hp1 : f p < 1) (hmin : ∀ m : ℕ, 0 < f 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 +lemma is_prime_of_minimal_nat_zero_lt_and_lt_one : p.Prime := by rw [← Nat.irreducible_iff_nat_prime] constructor -- Two goals: p is not a unit and any product giving p must contain a unit. · rw [Nat.isUnit_iff] @@ -177,14 +157,14 @@ lemma is_prime_of_minimal_nat_zero_lt_mulRingNorm_lt_one : p.Prime := by by_contra! con obtain ⟨ha₁, hb₁⟩ := con obtain ⟨ha₀, hb₀⟩ : a ≠ 0 ∧ b ≠ 0 := by - refine mul_ne_zero_iff.1 fun h ↦ ?_ + refine mul_ne_zero_iff.mp fun h ↦ ?_ rwa [h, Nat.cast_zero, map_zero, lt_self_iff_false] at hp0 have hap : a < a * b := lt_mul_of_one_lt_right (by omega) (by omega) have hbp : b < a * b := lt_mul_of_one_lt_left (by omega) (by omega) have ha := - le_of_not_lt <| not_and.1 ((hmin a).mt hap.not_le) (map_pos_of_ne_zero f (mod_cast ha₀)) + le_of_not_lt <| not_and.mp ((hmin a).mt hap.not_le) (map_pos_of_ne_zero f (mod_cast ha₀)) have hb := - le_of_not_lt <| not_and.1 ((hmin b).mt hbp.not_le) (map_pos_of_ne_zero f (mod_cast hb₀)) + le_of_not_lt <| not_and.mp ((hmin b).mt hbp.not_le) (map_pos_of_ne_zero f (mod_cast hb₀)) rw [Nat.cast_mul, map_mul] at hp1 exact ((one_le_mul_of_one_le_of_one_le ha hb).trans_lt hp1).false @@ -194,15 +174,14 @@ 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 +lemma eq_one_of_not_dvd {m : ℕ} (hpm : ¬ p ∣ m) : f m = 1 := by apply le_antisymm (bdd m) by_contra! hm set M := f p ⊔ f m with hM set k := Nat.ceil (M.logb (1 / 2)) + 1 with hk - obtain ⟨a, b, bezout⟩ : IsCoprime (p ^ k : ℤ) (m ^ k) := by - apply IsCoprime.pow (Nat.Coprime.isCoprime _) - exact (Nat.Prime.coprime_iff_not_dvd - (is_prime_of_minimal_nat_zero_lt_mulRingNorm_lt_one hp0 hp1 hmin)).2 hpm + obtain ⟨a, b, bezout⟩ : IsCoprime (p ^ k : ℤ) (m ^ k) := + is_prime_of_minimal_nat_zero_lt_and_lt_one hp0 hp1 hmin + |>.coprime_iff_not_dvd |>.mpr hpm |>.isCoprime |>.pow have le_half {x} (hx0 : 0 < x) (hx1 : x < 1) (hxM : x ≤ M) : x ^ k < 1 / 2 := by calc x ^ k = x ^ (k : ℝ) := (rpow_natCast x k).symm @@ -210,136 +189,123 @@ lemma mulRingNorm_eq_one_of_not_dvd {m : ℕ} (hpm : ¬ p ∣ m) : f m = 1 := by apply rpow_lt_rpow_of_exponent_gt hx0 hx1 rw [hk] push_cast - exact lt_add_of_le_of_pos (Nat.le_ceil (M.logb (1 / 2))) zero_lt_one + exact lt_add_of_le_of_pos (Nat.le_ceil _) zero_lt_one _ ≤ x ^ x.logb (1 / 2) := by - apply rpow_le_rpow_of_exponent_ge hx0 (le_of_lt hx1) + apply rpow_le_rpow_of_exponent_ge hx0 hx1.le simp only [one_div, ← log_div_log, log_inv, neg_div, ← div_neg, hM] gcongr simp only [Left.neg_pos_iff] - exact log_neg (lt_sup_iff.2 <| Or.inl hp0) (sup_lt_iff.2 ⟨hp1, hm⟩) - _ = 1 / 2 := rpow_logb hx0 (ne_of_lt hx1) one_half_pos + exact log_neg (lt_sup_iff.mpr <| .inl hp0) (sup_lt_iff.mpr ⟨hp1, hm⟩) + _ = 1 / 2 := rpow_logb hx0 hx1.ne one_half_pos apply lt_irrefl (1 : ℝ) calc 1 = f 1 := (map_one f).symm _ = f (a * p ^ k + b * m ^ k) := by rw_mod_cast [bezout]; norm_cast - _ ≤ f (a * p ^ k) + f (b * m ^ k) := f.add_le' (a * p ^ k) (b * m ^ k) + _ ≤ f (a * p ^ k) + f (b * m ^ k) := f.add_le' .. _ ≤ 1 * (f p) ^ k + 1 * (f m) ^ k := by - simp only [map_mul, map_pow, le_refl] + simp only [map_mul, map_pow] gcongr - all_goals rw [← MulRingNorm.apply_natAbs_eq]; apply bdd + all_goals rw [← apply_natAbs_eq]; apply bdd _ = (f p) ^ k + (f m) ^ k := by simp only [one_mul] _ < 1 := by - have hm₀ : 0 < f m := - map_pos_of_ne_zero _ <| Nat.cast_ne_zero.2 fun H ↦ hpm <| H ▸ dvd_zero p + have hm₀ : 0 < f m := f.pos <| Nat.cast_ne_zero.mpr fun H ↦ hpm <| H ▸ dvd_zero p linarith only [le_half hp0 hp1 le_sup_left, le_half hm₀ hm le_sup_right] --- ## Step 4: f p = p ^ (- t) for some positive real t +-- ## 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 - refine ⟨- logb p (f p), Left.neg_pos_iff.2 <| logb_neg (mod_cast pprime.one_lt) hp0 hp1, ?_⟩ +lemma exists_pos_eq_pow_neg : ∃ t : ℝ, 0 < t ∧ f p = p ^ (-t) := by + have pprime := is_prime_of_minimal_nat_zero_lt_and_lt_one hp0 hp1 hmin + refine ⟨- logb p (f p), Left.neg_pos_iff.mpr <| logb_neg (mod_cast pprime.one_lt) hp0 hp1, ?_⟩ rw [neg_neg] - refine (rpow_logb (mod_cast pprime.pos) ?_ hp0).symm - simp only [ne_eq, Nat.cast_eq_one,Nat.Prime.ne_one pprime, not_false_eq_true] + exact (rpow_logb (mod_cast pprime.pos) (mod_cast pprime.ne_one) hp0).symm -- ## 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, ∃ (_ : Fact (p.Prime)), MulRingNorm.equiv f (mulRingNorm_padic p) := by - obtain ⟨p, hfp, hmin⟩ := exists_minimal_nat_zero_lt_mulRingNorm_lt_one hf_nontriv bdd - have hprime := is_prime_of_minimal_nat_zero_lt_mulRingNorm_lt_one hfp.1 hfp.2 hmin - have hprime_fact : Fact (p.Prime) := ⟨hprime⟩ - obtain ⟨t, h⟩ := exists_pos_mulRingNorm_eq_pow_neg hfp.1 hfp.2 hmin +theorem equiv_padic_of_bounded : + ∃! p, ∃ (_ : Fact p.Prime), f ≈ (padic p) := by + obtain ⟨p, hfp, hmin⟩ := exists_minimal_nat_zero_lt_and_lt_one hf_nontriv bdd + have hprime := is_prime_of_minimal_nat_zero_lt_and_lt_one hfp.1 hfp.2 hmin + have hprime_fact : Fact p.Prime := ⟨hprime⟩ + obtain ⟨t, h⟩ := exists_pos_eq_pow_neg hfp.1 hfp.2 hmin simp_rw [← equiv_on_nat_iff_equiv] - use p - constructor -- 2 goals: MulRingNorm.equiv f (mulRingNorm_padic p) and p is unique. - · use hprime_fact - refine ⟨t⁻¹, by simp only [inv_pos, h.1], fun n ↦ ?_⟩ - have ht : t⁻¹ ≠ 0 := inv_ne_zero h.1.ne' - rcases eq_or_ne n 0 with rfl | hn -- Separate cases n=0 and n ≠ 0 - · simp only [Nat.cast_zero, map_zero, ne_eq, ht, not_false_eq_true, zero_rpow] + refine ⟨p, ⟨hprime_fact, t⁻¹, inv_pos_of_pos h.1, fun n ↦ ?_⟩, fun q ⟨hq_prime, h_equiv⟩ ↦ ?_⟩ + · have ht : t⁻¹ ≠ 0 := inv_ne_zero h.1.ne' + rcases eq_or_ne n 0 with rfl | hn -- Separate cases n = 0 and n ≠ 0 + · simp [ht] · /- Any natural number can be written as a power of p times a natural number not divisible by p -/ rcases Nat.exists_eq_pow_mul_and_not_dvd hn p hprime.ne_one with ⟨e, m, hpm, rfl⟩ - simp only [Nat.cast_mul, Nat.cast_pow, map_mul, map_pow, mulRingNorm_eq_padic_norm, - padicNorm.padicNorm_p_of_prime, Rat.cast_inv, Rat.cast_natCast, inv_pow, - mulRingNorm_eq_one_of_not_dvd bdd hfp.1 hfp.2 hmin hpm, h.2] + simp only [Nat.cast_mul, Nat.cast_pow, map_mul, map_pow, h.2, + eq_one_of_not_dvd bdd hfp.1 hfp.2 hmin hpm, padic_eq_padicNorm, + padicNorm.padicNorm_p_of_prime, cast_inv, cast_natCast, inv_pow] rw [← padicNorm.nat_eq_one_iff] at hpm - simp only [← rpow_natCast, p.cast_nonneg, ← rpow_mul, mul_one, ← rpow_neg, hpm, cast_one] + simp only [← rpow_natCast, p.cast_nonneg, ← rpow_mul, neg_mul, mul_one, ← rpow_neg, hpm, + cast_one] congr field_simp [h.1.ne'] - ring - · intro q ⟨hq_prime, h_equiv⟩ - by_contra! hne - apply Prime.ne_one (Nat.Prime.prime (Fact.elim hq_prime)) - rw [ne_comm, ← Nat.coprime_primes hprime (Fact.elim hq_prime), - Nat.Prime.coprime_iff_not_dvd hprime] at hne + · by_contra! hne + apply hq_prime.elim.prime.ne_one + rw [ne_comm, ← Nat.coprime_primes hprime hq_prime.elim, hprime.coprime_iff_not_dvd] at hne rcases h_equiv with ⟨c, _, h_eq⟩ have h_eq' := h_eq q - simp only [mulRingNorm_eq_one_of_not_dvd bdd hfp.1 hfp.2 hmin hne, one_rpow, - mulRingNorm_eq_padic_norm, padicNorm.padicNorm_p_of_prime, cast_inv, cast_natCast, eq_comm, - inv_eq_one] at h_eq' - norm_cast at h_eq' + simp only [eq_one_of_not_dvd bdd hfp.1 hfp.2 hmin hne, one_rpow, padic_eq_padicNorm, + padicNorm.padicNorm_p_of_prime, cast_inv, cast_natCast, eq_comm, inv_eq_one] at h_eq' + exact_mod_cast h_eq' end Non_archimedean section Archimedean -/-! ## Archimedean case +/-! +### Archimedean case -Every unbounded absolute value is equivalent to the standard absolute value +Every unbounded absolute value on `ℚ` is equivalent to the standard absolute value. -/ -/-- The usual absolute value on ℚ. -/ -def mulRingNorm_real : MulRingNorm ℚ := -{ toFun := fun x : ℚ ↦ |x| - map_zero' := by simp only [Rat.cast_zero, abs_zero] - add_le' := norm_add_le - neg' := norm_neg - eq_zero_of_map_eq_zero' := by simp only [abs_eq_zero, Rat.cast_eq_zero, imp_self, forall_const] - map_one' := by simp only [Rat.cast_one, abs_one] - map_mul' := by - simp only [Rat.cast_mul] - exact_mod_cast abs_mul -} - -@[simp] lemma mul_ring_norm_eq_abs (r : ℚ) : mulRingNorm_real r = |r| := by - simp only [Rat.cast_abs] - rfl +/-- The standard absolute value on `ℚ`. We name it `real` because it corresponds to the +unique real place of `ℚ`. -/ +def real : AbsoluteValue ℚ ℝ where + toFun x := |x| + map_mul' x y := by simpa using abs_mul (x : ℝ) (y : ℝ) + nonneg' x := by simp + eq_zero' x := by simp + add_le' x y := by simpa using abs_add_le (x : ℝ) (y : ℝ) + +@[simp] lemma real_eq_abs (r : ℚ) : real r = |r| := + (cast_abs r).symm -- ## Preliminary result -/-- Given an two integers `n, m` with `m > 1` the mulRingNorm of `n` is bounded by +/-- Given any two integers `n`, `m` with `m > 1`, the absolute value of `n` is bounded by `m + m * f m + m * (f m) ^ 2 + ... + m * (f m) ^ d` where `d` is the number of digits of the expansion of `n` in base `m`. -/ -lemma mulRingNorm_apply_le_sum_digits (n : ℕ) {m : ℕ} (hm : 1 < m) : +lemma apply_le_sum_digits (n : ℕ) {m : ℕ} (hm : 1 < m) : f n ≤ ((Nat.digits m n).mapIdx fun i _ ↦ m * (f m) ^ i).sum := by set L := Nat.digits m n set L' : List ℚ := List.map Nat.cast (L.mapIdx fun i a ↦ (a * m ^ i)) with hL' -- If `c` is a digit in the expansion of `n` in base `m`, then `f c` is less than `m`. have hcoef {c : ℕ} (hc : c ∈ Nat.digits m n) : f c < m := - lt_of_le_of_lt (MulRingNorm_nat_le_nat c f) (mod_cast Nat.digits_lt_base hm hc) + lt_of_le_of_lt (f.apply_nat_le_self c) (mod_cast Nat.digits_lt_base hm hc) calc f n = f ((Nat.ofDigits m L : ℕ) : ℚ) := by rw [Nat.ofDigits_digits m n] - _ = f (L'.sum) := by rw [Nat.ofDigits_eq_sum_mapIdx]; norm_cast - _ ≤ (L'.map f).sum := mulRingNorm_sum_le_sum_mulRingNorm L' f + _ = f L'.sum := by rw [Nat.ofDigits_eq_sum_mapIdx]; norm_cast + _ ≤ (L'.map f).sum := listSum_le f L' _ ≤ (L.mapIdx fun i _ ↦ m * (f m) ^ i).sum := ?_ simp only [hL', List.mapIdx_eq_enum_map, List.map_map] - apply List.sum_le_sum - rintro ⟨i, a⟩ hia - dsimp [Function.uncurry] + refine List.sum_le_sum fun ⟨i, a⟩ hia ↦ ?_ + dsimp only [Function.comp_apply, Function.uncurry_apply_pair] replace hia := List.mem_enumFrom hia push_cast rw [map_mul, map_pow] - gcongr + refine mul_le_mul_of_nonneg_right ?_ <| pow_nonneg (f.nonneg _) i simp only [zero_le, zero_add, tsub_zero, true_and] at hia exact (hcoef (List.mem_iff_get.mpr ⟨⟨i, hia.1⟩, hia.2.symm⟩)).le --- ## Step 1: if f is a MulRingNorm and f n > 1 for some natural n, then f n > 1 for all n ≥ 2 +-- ## Step 1: if f is an AbsoluteValue and f n > 1 for some natural n, then f n > 1 for all n ≥ 2 /-- If `f n > 1` for some `n` then `f n > 1` for all `n ≥ 2` -/ lemma one_lt_of_not_bounded (notbdd : ¬ ∀ n : ℕ, f n ≤ 1) {n₀ : ℕ} (hn₀ : 1 < n₀) : 1 < f n₀ := by @@ -349,11 +315,10 @@ lemma one_lt_of_not_bounded (notbdd : ¬ ∀ n : ℕ, f n ≤ 1) {n₀ : ℕ} (h /- L is the string of digits of `n` in the base `n₀`-/ set L := Nat.digits n₀ m calc - f m ≤ (L.mapIdx fun i _ ↦ n₀ * f n₀ ^ i).sum := mulRingNorm_apply_le_sum_digits m hn₀ + f m ≤ (L.mapIdx fun i _ ↦ n₀ * f n₀ ^ i).sum := apply_le_sum_digits m hn₀ _ ≤ (L.mapIdx fun _ _ ↦ (n₀ : ℝ)).sum := by simp only [List.mapIdx_eq_enum_map, List.map_map] - apply List.sum_le_sum - rintro ⟨i, a⟩ _ + refine List.sum_le_sum fun ⟨i, a⟩ _ ↦ ?_ simp only [Function.comp_apply, Function.uncurry_apply_pair] exact mul_le_of_le_of_le_one' (mod_cast le_refl n₀) (pow_le_one₀ (by positivity) h) (by positivity) (by positivity) @@ -366,157 +331,151 @@ lemma one_lt_of_not_bounded (notbdd : ¬ ∀ n : ℕ, f n ≤ 1) {n₀ : ℕ} (h _ ≤ n₀ * (logb n₀ m + 1) := by gcongr; exact natLog_le_logb .. -- For h_ineq2 we need to exclude the case n = 0. rcases eq_or_ne n 0 with rfl | h₀ - · simp only [CharP.cast_eq_zero, map_zero, zero_le_one] + · simp have h_ineq2 (k : ℕ) (hk : 0 < k) : f n ≤ (n₀ * (logb n₀ n + 1)) ^ (k : ℝ)⁻¹ * k ^ (k : ℝ)⁻¹ := by - have : 0 ≤ logb n₀ n := logb_nonneg (one_lt_cast.2 hn₀) (mod_cast Nat.one_le_of_lt h₀.bot_lt) + have : 0 ≤ logb n₀ n := logb_nonneg (one_lt_cast.mpr hn₀) (mod_cast Nat.one_le_of_lt h₀.bot_lt) calc f n = (f ↑(n ^ k)) ^ (k : ℝ)⁻¹ := by rw [Nat.cast_pow, map_pow, ← rpow_natCast, rpow_rpow_inv (by positivity) (by positivity)] - _ ≤ (n₀ * (logb n₀ ↑(n ^ k) + 1)) ^ (k : ℝ)⁻¹ := by + _ ≤ (n₀ * (logb n₀ ↑(n ^ k) + 1)) ^ (k : ℝ)⁻¹ := by gcongr exact h_ineq1 <| one_le_pow₀ (one_le_iff_ne_zero.mpr h₀) - _ = (n₀ * (k * logb n₀ n + 1)) ^ (k : ℝ)⁻¹ := by + _ = (n₀ * (k * logb n₀ n + 1)) ^ (k : ℝ)⁻¹ := by rw [Nat.cast_pow, logb_pow] - _ ≤ (n₀ * ( k * logb n₀ n + k)) ^ (k : ℝ)⁻¹ := by + _ ≤ (n₀ * (k * logb n₀ n + k)) ^ (k : ℝ)⁻¹ := by gcongr exact one_le_cast.mpr hk _ = (n₀ * (logb n₀ n + 1)) ^ (k : ℝ)⁻¹ * k ^ (k : ℝ)⁻¹ := by rw [← mul_rpow (by positivity) (by positivity), mul_assoc, add_mul, one_mul, - mul_comm _ (k : ℝ)] + mul_comm _ (k : ℝ)] -- For 0 < logb n₀ n below we also need to exclude n = 1. rcases eq_or_ne n 1 with rfl | h₁ - · simp only [Nat.cast_one, map_one, le_refl] - refine le_of_tendsto_of_tendsto tendsto_const_nhds ?_ (eventually_atTop.2 ⟨1, h_ineq2⟩) + · simp + refine le_of_tendsto_of_tendsto tendsto_const_nhds ?_ (eventually_atTop.mpr ⟨1, h_ineq2⟩) nth_rw 2 [← mul_one 1] have : 0 < logb n₀ n := logb_pos (mod_cast hn₀) (by norm_cast; omega) - have hnlim : Tendsto (fun k : ℕ ↦ (n₀ * (logb n₀ n + 1)) ^ (k : ℝ)⁻¹) atTop (𝓝 1) := - tendsto_root_atTop_nhds_one (by positivity) - exact hnlim.mul tendsto_nat_rpow_div + exact (tendsto_const_rpow_inv (by positivity)).mul tendsto_nat_rpow_inv --- ## Step 2: given m,n ≥ 2 and |m|=m^s, |n|=n^t for s,t > 0, we have t ≤ s +-- ## Step 2: given m, n ≥ 2 and |m| = m^s, |n| = n^t for s, t > 0, we have t ≤ s -variable {m n : ℕ} (hm : 1 < m) (hn : 1 < n) (notbdd : ¬ ∀ (n : ℕ), f n ≤ 1) +variable {m n : ℕ} (hm : 1 < m) (hn : 1 < n) (notbdd : ¬ ∀ n : ℕ, f n ≤ 1) include hm notbdd in private lemma expr_pos : 0 < m * f m / (f m - 1) := by apply div_pos (mul_pos (mod_cast zero_lt_of_lt hm) - (map_pos_of_ne_zero f (mod_cast ne_zero_of_lt hm))) + (map_pos_of_ne_zero f (mod_cast ne_zero_of_lt hm))) linarith only [one_lt_of_not_bounded notbdd hm] include hn hm notbdd in private lemma param_upperbound {k : ℕ} (hk : k ≠ 0) : - f n ≤ (m * f m / (f m - 1)) ^ (k : ℝ)⁻¹ * (f m) ^ (logb m n) := by + f n ≤ (m * f m / (f m - 1)) ^ (k : ℝ)⁻¹ * f m ^ logb m n := by have h_ineq1 {m n : ℕ} (hm : 1 < m) (hn : 1 < n) : - f n ≤ (m * f m / (f m - 1)) * (f m) ^ (logb m n) := by + f n ≤ (m * f m / (f m - 1)) * f m ^ logb m n := by let d := Nat.log m n calc - f n ≤ ((Nat.digits m n).mapIdx fun i _ ↦ m * f m ^ i).sum := - mulRingNorm_apply_le_sum_digits n hm - _ = m * ((Nat.digits m n).mapIdx fun i _ ↦ (f m) ^ i).sum := list_mul_sum (m.digits n) (f m) m + f n ≤ ((Nat.digits m n).mapIdx fun i _ ↦ m * f m ^ i).sum := apply_le_sum_digits n hm + _ = m * ((Nat.digits m n).mapIdx fun i _ ↦ f m ^ i).sum := list_mul_sum (m.digits n) (f m) m _ = m * ((f m ^ (d + 1) - 1) / (f m - 1)) := by rw [list_geom _ (ne_of_gt (one_lt_of_not_bounded notbdd hm)), - (Nat.digits_len m n hm (not_eq_zero_of_lt hn)).symm] - _ ≤ m * ((f m ^ (d + 1))/(f m - 1)) := by + ← Nat.digits_len m n hm (not_eq_zero_of_lt hn)] + _ ≤ m * ((f m ^ (d + 1)) / (f m - 1)) := by gcongr · linarith only [one_lt_of_not_bounded notbdd hm] - · simp only [tsub_le_iff_right, le_add_iff_nonneg_right, zero_le_one] + · simp _ = ↑m * f ↑m / (f ↑m - 1) * f ↑m ^ d := by ring _ ≤ ↑m * f ↑m / (f ↑m - 1) * f ↑m ^ logb ↑m ↑n := by gcongr - · exact le_of_lt (expr_pos hm notbdd) - · rw [← Real.rpow_natCast, Real.rpow_le_rpow_left_iff (one_lt_of_not_bounded notbdd hm)] + · exact (expr_pos hm notbdd).le + · rw [← rpow_natCast, rpow_le_rpow_left_iff (one_lt_of_not_bounded notbdd hm)] exact natLog_le_logb n m - apply le_of_pow_le_pow_left₀ hk (mul_nonneg (rpow_nonneg - (le_of_lt (expr_pos hm notbdd)) (k : ℝ)⁻¹) (rpow_nonneg (apply_nonneg f ↑m) (logb m n))) - nth_rw 2 [← Real.rpow_natCast] - rw [mul_rpow (rpow_nonneg (le_of_lt (expr_pos hm notbdd)) (k : ℝ)⁻¹) - (rpow_nonneg (apply_nonneg f ↑m) (logb ↑m ↑n)), ← rpow_mul (le_of_lt (expr_pos hm notbdd)), - ← rpow_mul (apply_nonneg f ↑m), inv_mul_cancel₀ (mod_cast hk), rpow_one, mul_comm (logb ↑m ↑n)] + apply le_of_pow_le_pow_left₀ hk <| mul_nonneg (rpow_nonneg (expr_pos hm notbdd).le _) + (rpow_nonneg (apply_nonneg f ↑m) _) + nth_rewrite 2 [← rpow_natCast] + rw [mul_rpow (rpow_nonneg (expr_pos hm notbdd).le _) (rpow_nonneg (apply_nonneg f ↑m) _), + ← rpow_mul (expr_pos hm notbdd).le, ← rpow_mul (apply_nonneg f ↑m), + inv_mul_cancel₀ (mod_cast hk), rpow_one, mul_comm (logb ..)] calc - (f n) ^ k = f ↑(n ^ k) := by simp only [Nat.cast_pow, map_pow] - _ ≤ (m * f m / (f m - 1)) * (f m) ^ (logb m ↑(n ^ k)) := h_ineq1 hm (Nat.one_lt_pow hk hn) - _ = (m * f m / (f m - 1)) * (f m) ^ (k * logb m n) := by - rw [Nat.cast_pow, Real.logb_pow] + (f n) ^ k = f ↑(n ^ k) := by simp + _ ≤ (m * f m / (f m - 1)) * f m ^ logb m ↑(n ^ k) := h_ineq1 hm (Nat.one_lt_pow hk hn) + _ = (m * f m / (f m - 1)) * f m ^ (k * logb m n) := by rw [Nat.cast_pow, logb_pow] include hm hn notbdd in /-- Given two natural numbers `n, m` greater than 1 we have `f n ≤ f m ^ logb m n`. -/ -lemma mulRingNorm_le_mulRingNorm_pow_log : f n ≤ f m ^ logb m n := by - have : Tendsto (fun k : ℕ ↦ (m * f m / (f m - 1)) ^ (k : ℝ)⁻¹ * (f m) ^ (logb m n)) - atTop (𝓝 ((f m) ^ (logb m n))) := by +lemma le_pow_log : f n ≤ f m ^ logb m n := by + have : Tendsto (fun k : ℕ ↦ (m * f m / (f m - 1)) ^ (k : ℝ)⁻¹ * f m ^ logb m n) + atTop (𝓝 (f m ^ logb m n)) := by nth_rw 2 [← one_mul (f ↑m ^ logb ↑m ↑n)] - exact Tendsto.mul_const _ (tendsto_root_atTop_nhds_one (expr_pos hm notbdd)) - exact le_of_tendsto_of_tendsto (tendsto_const_nhds (x:= f ↑n)) this ((eventually_atTop.2 ⟨2, - fun b hb ↦ param_upperbound hm hn notbdd (not_eq_zero_of_lt hb)⟩)) + exact (tendsto_const_rpow_inv (expr_pos hm notbdd)).mul_const _ + exact le_of_tendsto_of_tendsto (tendsto_const_nhds (x:= f ↑n)) this <| + eventually_atTop.mpr ⟨2, fun b hb ↦ param_upperbound hm hn notbdd (not_eq_zero_of_lt hb)⟩ include hm hn notbdd in -/-- Given `m,n ≥ 2` and `f m = m ^ s`, `f n = n ^ t` for `s, t > 0`, we have `t ≤ s`. -/ -lemma le_of_mulRingNorm_eq {s t : ℝ} (hfm : f m = m ^ s) (hfn : f n = n ^ t) : t ≤ s := by - have hmn : f n ≤ f m ^ Real.logb m n := mulRingNorm_le_mulRingNorm_pow_log hm hn notbdd - rw [← Real.rpow_le_rpow_left_iff (x:=n) (mod_cast hn), ← hfn] - apply le_trans hmn - rw [hfm, ← Real.rpow_mul (Nat.cast_nonneg m), mul_comm, Real.rpow_mul (Nat.cast_nonneg m), - Real.rpow_logb (mod_cast zero_lt_of_lt hm) (mod_cast Nat.ne_of_lt' hm) - (mod_cast zero_lt_of_lt hn)] +/-- Given `m, n ≥ 2` and `f m = m ^ s`, `f n = n ^ t` for `s, t > 0`, we have `t ≤ s`. -/ +private lemma le_of_eq_pow {s t : ℝ} (hfm : f m = m ^ s) (hfn : f n = n ^ t) : t ≤ s := by + rw [← rpow_le_rpow_left_iff (x := n) (mod_cast hn), ← hfn] + apply le_trans <| le_pow_log hm hn notbdd + rw [hfm, ← rpow_mul (Nat.cast_nonneg m), mul_comm, rpow_mul (Nat.cast_nonneg m), + rpow_logb (mod_cast zero_lt_of_lt hm) (mod_cast hm.ne') (mod_cast zero_lt_of_lt hn)] include hm hn notbdd in -private lemma symmetric_roles {s t : ℝ} (hfm : f m = m ^ s) (hfn : f n = n ^ t) : s = t := - le_antisymm (le_of_mulRingNorm_eq hn hm notbdd hfn hfm) - (le_of_mulRingNorm_eq hm hn notbdd hfm hfn) +private lemma eq_of_eq_pow {s t : ℝ} (hfm : f m = m ^ s) (hfn : f n = n ^ t) : s = t := + le_antisymm (le_of_eq_pow hn hm notbdd hfn hfm) (le_of_eq_pow hm hn notbdd hfm hfn) -- ## Archimedean case: end goal include notbdd in /-- If `f` is not bounded and not trivial, then it is equivalent to the standard absolute value on `ℚ`. -/ -theorem mulRingNorm_equiv_standard_of_unbounded : MulRingNorm.equiv f mulRingNorm_real := by +theorem equiv_real_of_unbounded : f ≈ real := by obtain ⟨m, hm⟩ := Classical.exists_not_of_not_forall notbdd have oneltm : 1 < m := by - by_contra! - apply hm - replace this : m = 0 ∨ m = 1 := by omega - rcases this with (rfl | rfl) - all_goals simp only [CharP.cast_eq_zero, map_zero, zero_le_one, Nat.cast_one, map_one, le_refl] + contrapose! hm + rcases le_one_iff_eq_zero_or_eq_one.mp hm with rfl | rfl <;> simp rw [← equiv_on_nat_iff_equiv] - set s := Real.logb m (f m) with hs - use s⁻¹ - refine ⟨inv_pos.2 (Real.logb_pos (Nat.one_lt_cast.2 oneltm) - (one_lt_of_not_bounded notbdd oneltm)), ?_⟩ - intro n - by_cases h1 : n ≤ 1 - · by_cases h2 : n = 1 - · simp only [h2, Nat.cast_one, map_one, one_rpow, abs_one, cast_one] - · have : n = 0 := by omega - rw [this, hs] - simp only [CharP.cast_eq_zero, map_zero] - rw [Real.rpow_eq_zero le_rfl] + set s := logb m (f m) with hs + refine ⟨s⁻¹, + inv_pos.mpr (logb_pos (Nat.one_lt_cast.mpr oneltm) (one_lt_of_not_bounded notbdd oneltm)), + fun n ↦ ?_⟩ + rcases lt_trichotomy n 1 with h | rfl | h + · obtain rfl : n = 0 := by omega + have : (logb (↑m) (f ↑m))⁻¹ ≠ 0 := by simp only [ne_eq, inv_eq_zero, logb_eq_zero, Nat.cast_eq_zero, Nat.cast_eq_one, map_eq_zero, not_or] - push_neg - exact ⟨not_eq_zero_of_lt oneltm, Nat.ne_of_lt' oneltm, mod_cast (fun a ↦ a), + exact ⟨not_eq_zero_of_lt oneltm, oneltm.ne', by norm_cast, not_eq_zero_of_lt oneltm, ne_of_not_le hm, by linarith only [apply_nonneg f ↑m]⟩ - · simp only [mul_ring_norm_eq_abs, abs_cast, Rat.cast_natCast] - rw [Real.rpow_inv_eq (apply_nonneg f ↑n) (Nat.cast_nonneg n) - (Real.logb_ne_zero_of_pos_of_ne_one (one_lt_cast.mpr oneltm) (by linarith only [hm]) + simp [hs, this] + · simp + · simp only [real_eq_abs, abs_cast, Rat.cast_natCast] + rw [rpow_inv_eq (apply_nonneg f ↑n) (Nat.cast_nonneg n) + (logb_ne_zero_of_pos_of_ne_one (one_lt_cast.mpr oneltm) (by linarith only [hm]) (by linarith only [hm]))] - simp only [not_le] at h1 - have hfm : f m = m ^ s := by rw [Real.rpow_logb (mod_cast zero_lt_of_lt oneltm) - (mod_cast Nat.ne_of_lt' oneltm) (by linarith only [hm])] - have hfn : f n = n ^ (Real.logb n (f n)) := by - rw [Real.rpow_logb (mod_cast zero_lt_of_lt h1) (mod_cast Nat.ne_of_lt' h1) - (by apply map_pos_of_ne_zero; exact_mod_cast not_eq_zero_of_lt h1)] - rwa [← hs, symmetric_roles oneltm h1 notbdd hfm hfn] + have hfm : f m = m ^ s := by + rw [rpow_logb (mod_cast zero_lt_of_lt oneltm) (mod_cast oneltm.ne') (by linarith only [hm])] + have hfn : f n = n ^ logb n (f n) := by + rw [rpow_logb (mod_cast zero_lt_of_lt h) (mod_cast h.ne') + (by apply map_pos_of_ne_zero; exact_mod_cast not_eq_zero_of_lt h)] + rwa [← hs, eq_of_eq_pow oneltm h notbdd hfm hfn] end Archimedean -/-- **Ostrowski's Theorem** -/ -theorem mulRingNorm_equiv_standard_or_padic (f : MulRingNorm ℚ) (hf_nontriv : f ≠ 1) : - (MulRingNorm.equiv f mulRingNorm_real) ∨ - ∃! p, ∃ (_ : Fact (p.Prime)), MulRingNorm.equiv f (mulRingNorm_padic p) := by +/-! +### The main result +-/ + +/-- **Ostrowski's Theorem**: every absolute value (with values in `ℝ`) on `ℚ` is equivalent +to either the standard absolute value or a `p`-adic absolute value for a prime `p`. -/ +theorem equiv_real_or_padic (f : AbsoluteValue ℚ ℝ) (hf_nontriv : f ≠ .trivial) : + f ≈ real ∨ ∃! p, ∃ (_ : Fact p.Prime), f ≈ (padic p) := by by_cases bdd : ∀ n : ℕ, f n ≤ 1 - · right - exact mulRingNorm_equiv_padic_of_bounded hf_nontriv bdd - · left - exact mulRingNorm_equiv_standard_of_unbounded bdd + · exact .inr <| equiv_padic_of_bounded hf_nontriv bdd + · exact .inl <| equiv_real_of_unbounded bdd + +/-- The standard absolute value on `ℚ` is not equivalent to any `p`-adic absolute value. -/ +lemma not_real_equiv_padic (p : ℕ) [Fact p.Prime] : ¬ real ≈ (padic p) := by + rintro ⟨c, hc₀, hc⟩ + apply_fun (· 2) at hc + simp only [real_eq_abs, abs_ofNat, cast_ofNat] at hc + exact ((padic_le_one p 2).trans_lt <| one_lt_rpow one_lt_two hc₀).ne' hc -end Rat.MulRingNorm +end Rat.AbsoluteValue diff --git a/Mathlib/Topology/Algebra/Ring/Basic.lean b/Mathlib/Topology/Algebra/Ring/Basic.lean index b31686155cefd..f565a4012f388 100644 --- a/Mathlib/Topology/Algebra/Ring/Basic.lean +++ b/Mathlib/Topology/Algebra/Ring/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Johannes Hölzl -/ -import Mathlib.Algebra.Order.AbsoluteValue +import Mathlib.Algebra.Order.AbsoluteValue.Basic import Mathlib.Algebra.Ring.Prod import Mathlib.Algebra.Ring.Subring.Basic import Mathlib.Topology.Algebra.Group.Basic diff --git a/Mathlib/Topology/UniformSpace/AbsoluteValue.lean b/Mathlib/Topology/UniformSpace/AbsoluteValue.lean index b10faf84a84ef..7fac2dedb15f7 100644 --- a/Mathlib/Topology/UniformSpace/AbsoluteValue.lean +++ b/Mathlib/Topology/UniformSpace/AbsoluteValue.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot -/ -import Mathlib.Algebra.Order.AbsoluteValue +import Mathlib.Algebra.Order.AbsoluteValue.Basic import Mathlib.Algebra.Order.Field.Basic import Mathlib.Topology.UniformSpace.OfFun diff --git a/MathlibTest/ValuedCSP.lean b/MathlibTest/ValuedCSP.lean index 48b89612eacdc..e3384693d5102 100644 --- a/MathlibTest/ValuedCSP.lean +++ b/MathlibTest/ValuedCSP.lean @@ -1,4 +1,4 @@ -import Mathlib.Algebra.Order.AbsoluteValue +import Mathlib.Algebra.Order.AbsoluteValue.Basic import Mathlib.Algebra.Order.Ring.Rat import Mathlib.Combinatorics.Optimization.ValuedCSP import Mathlib.Data.Fin.Tuple.Curry From ec3b0067299cb6178b8ef0e3a134e76646f4a909 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Thu, 2 Jan 2025 22:00:59 +0000 Subject: [PATCH 097/189] chore(Algebra/Lie/Solvable): add instance IsSolvable R A for LieIdeal R L (#20395) --- Mathlib/Algebra/Lie/Solvable.lean | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Lie/Solvable.lean b/Mathlib/Algebra/Lie/Solvable.lean index 4a9632102afc4..a17e1accef60b 100644 --- a/Mathlib/Algebra/Lie/Solvable.lean +++ b/Mathlib/Algebra/Lie/Solvable.lean @@ -229,6 +229,9 @@ theorem Injective.lieAlgebra_isSolvable [h₁ : IsSolvable R L] (h₂ : Injectiv apply LieIdeal.bot_of_map_eq_bot h₂; rw [eq_bot_iff, ← hk] apply LieIdeal.derivedSeries_map_le +instance (A : LieIdeal R L) [IsSolvable R L] : IsSolvable R A := + A.incl_injective.lieAlgebra_isSolvable + theorem Surjective.lieAlgebra_isSolvable [h₁ : IsSolvable R L'] (h₂ : Surjective f) : IsSolvable R L := by obtain ⟨k, hk⟩ := id h₁ @@ -238,7 +241,7 @@ theorem Surjective.lieAlgebra_isSolvable [h₁ : IsSolvable R L'] (h₂ : Surjec end Function -theorem LieHom.isSolvable_range (f : L' →ₗ⁅R⁆ L) [LieAlgebra.IsSolvable R L'] : +instance LieHom.isSolvable_range (f : L' →ₗ⁅R⁆ L) [LieAlgebra.IsSolvable R L'] : LieAlgebra.IsSolvable R f.range := f.surjective_rangeRestrict.lieAlgebra_isSolvable From f36aaf3682d93c2791bee76ef69613ff8c88f0fa Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Thu, 2 Jan 2025 22:11:06 +0000 Subject: [PATCH 098/189] chore(Order/Atoms): Add IsAtom.bot_lt and IsCoatom.lt_top (#20396) --- Mathlib/Order/Atoms.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/Mathlib/Order/Atoms.lean b/Mathlib/Order/Atoms.lean index a9f5023c860eb..8b875265266cf 100644 --- a/Mathlib/Order/Atoms.lean +++ b/Mathlib/Order/Atoms.lean @@ -89,6 +89,9 @@ theorem IsAtom.lt_iff (h : IsAtom a) : x < a ↔ x = ⊥ := theorem IsAtom.le_iff (h : IsAtom a) : x ≤ a ↔ x = ⊥ ∨ x = a := by rw [le_iff_lt_or_eq, h.lt_iff] +lemma IsAtom.bot_lt (h : IsAtom a) : ⊥ < a := + h.lt_iff.mpr rfl + lemma IsAtom.le_iff_eq (ha : IsAtom a) (hb : b ≠ ⊥) : b ≤ a ↔ b = a := ha.le_iff.trans <| or_iff_right hb @@ -166,6 +169,9 @@ theorem IsCoatom.lt_iff (h : IsCoatom a) : a < x ↔ x = ⊤ := theorem IsCoatom.le_iff (h : IsCoatom a) : a ≤ x ↔ x = ⊤ ∨ x = a := h.dual.le_iff +lemma IsCoatom.lt_top (h : IsCoatom a) : a < ⊤ := + h.lt_iff.mpr rfl + lemma IsCoatom.le_iff_eq (ha : IsCoatom a) (hb : b ≠ ⊤) : a ≤ b ↔ b = a := ha.dual.le_iff_eq hb theorem IsCoatom.Ici_eq (h : IsCoatom a) : Set.Ici a = {⊤, a} := From ba6854b002dfa5265e40312c4c27a4f947b0c352 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 2 Jan 2025 23:22:49 +0000 Subject: [PATCH 099/189] chore(SetTheory/Game/Ordinal): make `Ordinal.toGame` an order embedding (#19417) The original reasoning against doing this was that it would disable dot notation, which is no longer an issue in the latest version of Lean. --- Mathlib/SetTheory/Game/Birthday.lean | 2 +- Mathlib/SetTheory/Game/Ordinal.lean | 37 ++++++++++++++-------------- 2 files changed, 20 insertions(+), 19 deletions(-) diff --git a/Mathlib/SetTheory/Game/Birthday.lean b/Mathlib/SetTheory/Game/Birthday.lean index 8be4a422d4485..c0da2ca0c18a8 100644 --- a/Mathlib/SetTheory/Game/Birthday.lean +++ b/Mathlib/SetTheory/Game/Birthday.lean @@ -207,7 +207,7 @@ theorem birthday_ordinalToGame (o : Ordinal) : birthday o.toGame = o := by apply birthday_quot_le_pGameBirthday · let ⟨x, hx₁, hx₂⟩ := birthday_eq_pGameBirthday o.toGame rw [← hx₂, ← toPGame_le_iff] - rw [← PGame.equiv_iff_game_eq] at hx₁ + rw [← mk_toPGame, ← PGame.equiv_iff_game_eq] at hx₁ exact hx₁.2.trans (PGame.le_birthday x) @[simp, norm_cast] diff --git a/Mathlib/SetTheory/Game/Ordinal.lean b/Mathlib/SetTheory/Game/Ordinal.lean index e770e5e3afd5d..855a0f7f2db00 100644 --- a/Mathlib/SetTheory/Game/Ordinal.lean +++ b/Mathlib/SetTheory/Game/Ordinal.lean @@ -157,15 +157,18 @@ noncomputable def toPGameEmbedding : Ordinal.{u} ↪o PGame.{u} where map_rel_iff' := @toPGame_le_iff /-- Converts an ordinal into the corresponding game. -/ -noncomputable abbrev toGame (o : Ordinal) : Game := ⟦o.toPGame⟧ - -/-- The order embedding version of `toGame`. -/ -@[simps] -noncomputable def toGameEmbedding : Ordinal.{u} ↪o Game.{u} where - toFun := Ordinal.toGame +noncomputable def toGame : Ordinal.{u} ↪o Game.{u} where + toFun o := ⟦o.toPGame⟧ inj' a b := by simpa [AntisymmRel] using le_antisymm map_rel_iff' := toPGame_le_iff +@[simp] +theorem mk_toPGame (o : Ordinal) : ⟦o.toPGame⟧ = o.toGame := + rfl + +@[deprecated toGame (since := "2024-11-23")] +alias toGameEmbedding := toGame + @[simp] theorem toGame_zero : toGame 0 = 0 := game_eq toPGame_zero @@ -175,22 +178,20 @@ theorem toGame_one : toGame 1 = 1 := game_eq toPGame_one theorem toGame_injective : Function.Injective toGame := - toGameEmbedding.injective + toGame.injective @[simp] -theorem toGame_lf_iff {a b : Ordinal} : Game.LF (toGame a) (toGame b) ↔ a < b := +theorem toGame_lf_iff {a b : Ordinal} : Game.LF a.toGame b.toGame ↔ a < b := toPGame_lf_iff -@[simp] -theorem toGame_le_iff {a b : Ordinal} : toGame a ≤ toGame b ↔ a ≤ b := +theorem toGame_le_iff {a b : Ordinal} : a.toGame ≤ b.toGame ↔ a ≤ b := toPGame_le_iff -@[simp] -theorem toGame_lt_iff {a b : Ordinal} : toGame a < toGame b ↔ a < b := +theorem toGame_lt_iff {a b : Ordinal} : a.toGame < b.toGame ↔ a < b := toPGame_lt_iff -theorem toGame_inj {a b : Ordinal} : toGame a = toGame b ↔ a = b := - toGameEmbedding.inj +theorem toGame_inj {a b : Ordinal} : a.toGame = b.toGame ↔ a = b := + toGame.inj @[deprecated (since := "2024-12-29")] alias toGame_eq_iff := toGame_inj @@ -223,7 +224,7 @@ theorem toPGame_nmul (a b : Ordinal) : (a ⨳ b).toPGame ≈ a.toPGame * b.toPGa 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 _ _, + rw [← toPGame_le_iff, le_iff_game_le, mk_toPGame, mk_toPGame, 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⟩⟩ @@ -236,7 +237,7 @@ theorem toPGame_nmul (a b : Ordinal) : (a ⨳ b).toPGame ≈ a.toPGame * b.toPGa 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] + simp_rw [mk_toPGame, ← toGame_nadd] apply toPGame_lf (nmul_nadd_lt _ _) <;> exact toLeftMovesToPGame_symm_lt _ termination_by (a, b) @@ -244,7 +245,7 @@ termination_by (a, b) theorem toGame_nmul (a b : Ordinal) : (a ⨳ b).toGame = ⟦a.toPGame * b.toPGame⟧ := game_eq (toPGame_nmul a b) -@[simp, norm_cast] +@[simp] -- used to be a norm_cast lemma theorem toGame_natCast : ∀ n : ℕ, toGame n = n | 0 => Quot.sound (zeroToPGameRelabelling).equiv | n + 1 => by @@ -253,6 +254,6 @@ theorem toGame_natCast : ∀ n : ℕ, toGame n = n rfl theorem toPGame_natCast (n : ℕ) : toPGame n ≈ n := by - rw [PGame.equiv_iff_game_eq, ← toGame, toGame_natCast, quot_natCast] + rw [PGame.equiv_iff_game_eq, mk_toPGame, toGame_natCast, quot_natCast] end Ordinal From 7bbc345c235e9076bea77f543abd430e11a9ffda Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Miyahara=20K=C5=8D?= Date: Fri, 3 Jan 2025 00:04:11 +0000 Subject: [PATCH 100/189] docs(Mathlib/NumberTheory/Bernoulli): update doc (#20377) --- Mathlib/NumberTheory/Bernoulli.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/Bernoulli.lean b/Mathlib/NumberTheory/Bernoulli.lean index dbd30c72d4719..8e39a25fbcd57 100644 --- a/Mathlib/NumberTheory/Bernoulli.lean +++ b/Mathlib/NumberTheory/Bernoulli.lean @@ -19,11 +19,11 @@ a sequence of rational numbers. They show up in the formula for the sums of $k$t powers. They are related to the Taylor series expansions of $x/\tan(x)$ and of $\coth(x)$, and also show up in the values that the Riemann Zeta function takes both at both negative and positive integers (and hence in the -theory of modular forms). For example, if $1 \leq n$ is even then +theory of modular forms). For example, if $1 \leq n$ then $$\zeta(2n)=\sum_{t\geq1}t^{-2n}=(-1)^{n+1}\frac{(2\pi)^{2n}B_{2n}}{2(2n)!}.$$ -Note however that this result is not yet formalised in Lean. +This result is formalised in Lean: `riemannZeta_two_mul_nat`. The Bernoulli numbers can be formally defined using the power series From add4d4ae372578422ece834d30ede1fed02dfebf Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 3 Jan 2025 05:55:58 +0000 Subject: [PATCH 101/189] chore(LinearAlgebra/Span/Basic): move/add lemmas about span_singleton (#20397) --- Mathlib/LinearAlgebra/Span/Basic.lean | 49 ++++++++++++++++----------- 1 file changed, 29 insertions(+), 20 deletions(-) diff --git a/Mathlib/LinearAlgebra/Span/Basic.lean b/Mathlib/LinearAlgebra/Span/Basic.lean index a6c4efb340ac8..7ef6263250020 100644 --- a/Mathlib/LinearAlgebra/Span/Basic.lean +++ b/Mathlib/LinearAlgebra/Span/Basic.lean @@ -148,20 +148,6 @@ theorem coe_scott_continuous : OmegaCompletePartialOrder.ωScottContinuous.of_monotone_map_ωSup ⟨SetLike.coe_mono, coe_iSup_of_chain⟩ -theorem disjoint_span_singleton {K E : Type*} [DivisionRing K] [AddCommGroup E] [Module K E] - {s : Submodule K E} {x : E} : Disjoint s (K ∙ x) ↔ x ∈ s → x = 0 := by - refine disjoint_def.trans ⟨fun H hx => H x hx <| subset_span <| mem_singleton x, ?_⟩ - intro H y hy hyx - obtain ⟨c, rfl⟩ := mem_span_singleton.1 hyx - by_cases hc : c = 0 - · rw [hc, zero_smul] - · rw [s.smul_mem_iff hc] at hy - rw [H hy, smul_zero] - -theorem disjoint_span_singleton' {K E : Type*} [DivisionRing K] [AddCommGroup E] [Module K E] - {p : Submodule K E} {x : E} (x0 : x ≠ 0) : Disjoint p (K ∙ x) ↔ x ∉ p := - disjoint_span_singleton.trans ⟨fun h₁ h₂ => x0 (h₁ h₂), fun h₁ h₂ => (h₁ h₂).elim⟩ - variable (R S s) /-- If `R` is "smaller" ring than `S` then the span by `R` is smaller than the span by `S`. -/ @@ -447,13 +433,13 @@ end AddCommGroup section DivisionRing -variable [DivisionRing K] [AddCommGroup V] [Module K V] +variable [DivisionRing K] [AddCommGroup V] [Module K V] {s : Submodule K V} {x : V} -/-- There is no vector subspace between `p` and `(K ∙ x) ⊔ p`, `WCovBy` version. -/ -theorem wcovBy_span_singleton_sup (x : V) (p : Submodule K V) : WCovBy p ((K ∙ x) ⊔ p) := by +/-- There is no vector subspace between `s` and `(K ∙ x) ⊔ s`, `WCovBy` version. -/ +theorem wcovBy_span_singleton_sup (x : V) (s : Submodule K V) : WCovBy s ((K ∙ x) ⊔ s) := by refine ⟨le_sup_right, fun q hpq hqp ↦ hqp.not_le ?_⟩ rcases SetLike.exists_of_lt hpq with ⟨y, hyq, hyp⟩ - obtain ⟨c, z, hz, rfl⟩ : ∃ c : K, ∃ z ∈ p, c • x + z = y := by + obtain ⟨c, z, hz, rfl⟩ : ∃ c : K, ∃ z ∈ s, c • x + z = y := by simpa [mem_sup, mem_span_singleton] using hqp.le hyq rcases eq_or_ne c 0 with rfl | hc · simp [hz] at hyp @@ -461,10 +447,33 @@ theorem wcovBy_span_singleton_sup (x : V) (p : Submodule K V) : WCovBy p ((K ∙ rwa [q.add_mem_iff_left (hpq.le hz), q.smul_mem_iff hc] at hyq simp [hpq.le, this] -/-- There is no vector subspace between `p` and `(K ∙ x) ⊔ p`, `CovBy` version. -/ -theorem covBy_span_singleton_sup {x : V} {p : Submodule K V} (h : x ∉ p) : CovBy p ((K ∙ x) ⊔ p) := +/-- There is no vector subspace between `s` and `(K ∙ x) ⊔ s`, `CovBy` version. -/ +theorem covBy_span_singleton_sup {x : V} {s : Submodule K V} (h : x ∉ s) : CovBy s ((K ∙ x) ⊔ s) := ⟨by simpa, (wcovBy_span_singleton_sup _ _).2⟩ +theorem disjoint_span_singleton : Disjoint s (K ∙ x) ↔ x ∈ s → x = 0 := by + refine disjoint_def.trans ⟨fun H hx => H x hx <| subset_span <| mem_singleton x, ?_⟩ + intro H y hy hyx + obtain ⟨c, rfl⟩ := mem_span_singleton.1 hyx + by_cases hc : c = 0 + · rw [hc, zero_smul] + · rw [s.smul_mem_iff hc] at hy + rw [H hy, smul_zero] + +theorem disjoint_span_singleton' (x0 : x ≠ 0) : Disjoint s (K ∙ x) ↔ x ∉ s := + disjoint_span_singleton.trans ⟨fun h₁ h₂ => x0 (h₁ h₂), fun h₁ h₂ => (h₁ h₂).elim⟩ + +lemma disjoint_span_singleton_of_not_mem (hx : x ∉ s) : Disjoint s (K ∙ x) := by + rw [disjoint_span_singleton] + intro h + contradiction + +lemma isCompl_span_singleton_of_isCoatom_of_not_mem (hs : IsCoatom s) (hx : x ∉ s) : + IsCompl s (K ∙ x) := by + refine ⟨disjoint_span_singleton_of_not_mem hx, ?_⟩ + rw [← covBy_top_iff] at hs + simpa only [codisjoint_iff, sup_comm, not_lt_top_iff] using hs.2 (covBy_span_singleton_sup hx).1 + end DivisionRing end Submodule From d01dbf92ed75e06b60de7987b5818266f6a9c6d0 Mon Sep 17 00:00:00 2001 From: Artie Khovanov Date: Fri, 3 Jan 2025 07:04:03 +0000 Subject: [PATCH 102/189] chore(Algebra/Lie): rename theorems for consistency (#20408) Cleaning up issues raised in #20353 Moves: lieIdealSubalgebra -> LieIdeal.toLieSubalgebra LieIdeal.coe_toLieSubalgebra_toSubmodule -> LieIdeal.toLieSubalgebra_toSubmodule --- Mathlib/Algebra/Lie/Engel.lean | 2 +- Mathlib/Algebra/Lie/InvariantForm.lean | 2 +- Mathlib/Algebra/Lie/Semisimple/Basic.lean | 2 +- Mathlib/Algebra/Lie/Submodule.lean | 19 ++++++++++++------- Mathlib/Algebra/Lie/TraceForm.lean | 2 +- 5 files changed, 16 insertions(+), 11 deletions(-) diff --git a/Mathlib/Algebra/Lie/Engel.lean b/Mathlib/Algebra/Lie/Engel.lean index 6cab534fb7f10..4380e2dd3aef3 100644 --- a/Mathlib/Algebra/Lie/Engel.lean +++ b/Mathlib/Algebra/Lie/Engel.lean @@ -198,7 +198,7 @@ theorem LieAlgebra.exists_engelian_lieSubalgebra_of_lt_normalizer {K : LieSubalg obtain ⟨I, hI₁ : (I : LieSubalgebra R K') = LieSubalgebra.ofLe hKK'⟩ := LieSubalgebra.exists_nested_lieIdeal_ofLe_normalizer hKK' hK' have hI₂ : (R ∙ (⟨x, hxK'⟩ : K')) ⊔ (LieSubmodule.toSubmodule I) = ⊤ := by - rw [← LieIdeal.coe_toLieSubalgebra_toSubmodule R K' I, hI₁] + rw [← LieIdeal.toLieSubalgebra_toSubmodule R K' I, hI₁] apply Submodule.map_injective_of_injective (K' : Submodule R L).injective_subtype simp only [LieSubalgebra.coe_ofLe, Submodule.map_sup, Submodule.map_subtype_range_inclusion, Submodule.map_top, Submodule.range_subtype] diff --git a/Mathlib/Algebra/Lie/InvariantForm.lean b/Mathlib/Algebra/Lie/InvariantForm.lean index 0722c41605ea8..2c422bdc8e5eb 100644 --- a/Mathlib/Algebra/Lie/InvariantForm.lean +++ b/Mathlib/Algebra/Lie/InvariantForm.lean @@ -150,7 +150,7 @@ lemma restrict_nondegenerate (I : LieIdeal K L) (hI : IsAtom I) : lemma restrict_orthogonal_nondegenerate (I : LieIdeal K L) (hI : IsAtom I) : (Φ.restrict (orthogonal Φ hΦ_inv I)).Nondegenerate := by rw [LinearMap.BilinForm.restrict_nondegenerate_iff_isCompl_orthogonal hΦ_refl] - simp only [LieIdeal.coe_toLieSubalgebra_toSubmodule, orthogonal_toSubmodule, + simp only [LieIdeal.toLieSubalgebra_toSubmodule, orthogonal_toSubmodule, LinearMap.BilinForm.orthogonal_orthogonal hΦ_nondeg hΦ_refl] exact (orthogonal_isCompl_toSubmodule Φ hΦ_nondeg hΦ_inv hΦ_refl hL I hI).symm diff --git a/Mathlib/Algebra/Lie/Semisimple/Basic.lean b/Mathlib/Algebra/Lie/Semisimple/Basic.lean index 2bcaab54edd7e..16117bb7c15bf 100644 --- a/Mathlib/Algebra/Lie/Semisimple/Basic.lean +++ b/Mathlib/Algebra/Lie/Semisimple/Basic.lean @@ -159,7 +159,7 @@ lemma isSimple_of_isAtom (I : LieIdeal R L) (hI : IsAtom I) : IsSimple R I where intro x hx suffices x ∈ J → x = 0 from this hx have := @this x.1 - simp only [LieIdeal.incl_coe, LieIdeal.coe_toLieSubalgebra_toSubmodule, + simp only [LieIdeal.incl_coe, LieIdeal.toLieSubalgebra_toSubmodule, LieSubmodule.mem_mk_iff', Submodule.mem_map, LieSubmodule.mem_toSubmodule, Subtype.exists, LieSubmodule.mem_bot, ZeroMemClass.coe_eq_zero, forall_exists_index, and_imp, J'] at this exact fun _ ↦ this (↑x) x.property hx rfl diff --git a/Mathlib/Algebra/Lie/Submodule.lean b/Mathlib/Algebra/Lie/Submodule.lean index 5b057903a082c..f72a8112bb070 100644 --- a/Mathlib/Algebra/Lie/Submodule.lean +++ b/Mathlib/Algebra/Lie/Submodule.lean @@ -224,11 +224,13 @@ theorem lie_mem_left (I : LieIdeal R L) (x y : L) (h : x ∈ I) : ⁅x, y⁆ ∈ rw [← lie_skew, ← neg_lie]; apply lie_mem_right; assumption /-- An ideal of a Lie algebra is a Lie subalgebra. -/ -def lieIdealSubalgebra (I : LieIdeal R L) : LieSubalgebra R L := +def LieIdeal.toLieSubalgebra (I : LieIdeal R L) : LieSubalgebra R L := { I.toSubmodule with lie_mem' := by intro x y _ hy; apply lie_mem_right; exact hy } +@[deprecated (since := "2025-01-02")] alias lieIdealSubalgebra := LieIdeal.toLieSubalgebra + instance : Coe (LieIdeal R L) (LieSubalgebra R L) := - ⟨lieIdealSubalgebra R L⟩ + ⟨LieIdeal.toLieSubalgebra R L⟩ @[simp] theorem LieIdeal.coe_toLieSubalgebra (I : LieIdeal R L) : ((I : LieSubalgebra R L) : Set L) = I := @@ -238,12 +240,15 @@ theorem LieIdeal.coe_toLieSubalgebra (I : LieIdeal R L) : ((I : LieSubalgebra R alias LieIdeal.coe_toSubalgebra := LieIdeal.coe_toLieSubalgebra @[simp] -theorem LieIdeal.coe_toLieSubalgebra_toSubmodule (I : LieIdeal R L) : +theorem LieIdeal.toLieSubalgebra_toSubmodule (I : LieIdeal R L) : ((I : LieSubalgebra R L) : Submodule R L) = LieSubmodule.toSubmodule I := rfl +@[deprecated (since := "2025-01-02")] +alias LieIdeal.coe_toLieSubalgebra_toSubmodule := LieIdeal.toLieSubalgebra_toSubmodule + @[deprecated (since := "2024-12-30")] -alias LieIdeal.coe_to_lieSubalgebra_to_submodule := LieIdeal.coe_toLieSubalgebra_toSubmodule +alias LieIdeal.coe_to_lieSubalgebra_to_submodule := LieIdeal.toLieSubalgebra_toSubmodule /-- An ideal of `L` is a Lie subalgebra of `L`, so it is a Lie ring. -/ instance LieIdeal.lieRing (I : LieIdeal R L) : LieRing I := @@ -300,7 +305,7 @@ theorem mem_toLieSubmodule (x : L) : x ∈ K.toLieSubmodule ↔ x ∈ K := theorem exists_lieIdeal_coe_eq_iff : (∃ I : LieIdeal R L, ↑I = K) ↔ ∀ x y : L, y ∈ K → ⁅x, y⁆ ∈ K := by - simp only [← toSubmodule_inj, LieIdeal.coe_toLieSubalgebra_toSubmodule, + simp only [← toSubmodule_inj, LieIdeal.toLieSubalgebra_toSubmodule, Submodule.exists_lieSubmodule_coe_eq_iff L] exact Iff.rfl @@ -1099,7 +1104,7 @@ theorem IsIdealMorphism.eq (hf : f.IsIdealMorphism) : f.idealRange = f.range := theorem isIdealMorphism_iff : f.IsIdealMorphism ↔ ∀ (x : L') (y : L), ∃ z : L, ⁅x, f y⁆ = f z := by simp only [isIdealMorphism_def, idealRange_eq_lieSpan_range, ← LieSubalgebra.toSubmodule_inj, ← f.range.coe_toSubmodule, - LieIdeal.coe_toLieSubalgebra_toSubmodule, LieSubmodule.coe_lieSpan_submodule_eq_iff, + LieIdeal.toLieSubalgebra_toSubmodule, LieSubmodule.coe_lieSpan_submodule_eq_iff, LieSubalgebra.mem_toSubmodule, mem_range, exists_imp, Submodule.exists_lieSubmodule_coe_eq_iff] constructor @@ -1288,7 +1293,7 @@ theorem ker_incl : I.incl.ker = ⊥ := by ext; simp @[simp] theorem incl_idealRange : I.incl.idealRange = I := by rw [LieHom.idealRange_eq_lieSpan_range, ← LieSubalgebra.coe_toSubmodule, ← - LieSubmodule.toSubmodule_inj, incl_range, coe_toLieSubalgebra_toSubmodule, + LieSubmodule.toSubmodule_inj, incl_range, toLieSubalgebra_toSubmodule, LieSubmodule.coe_lieSpan_submodule_eq_iff] use I diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean index 068aae6074787..40e83c9dbe56d 100644 --- a/Mathlib/Algebra/Lie/TraceForm.lean +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -287,7 +287,7 @@ lemma lowerCentralSeries_one_inf_center_le_ker_traceForm [Module.Free R M] [Modu 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_toLieSubalgebra_toSubmodule, LieSubmodule.toSubmodule_eq_bot, h] + LieIdeal.toLieSubalgebra_toSubmodule, LieSubmodule.toSubmodule_eq_bot, h] using lowerCentralSeries_one_inf_center_le_ker_traceForm R L M end LieModule From b2158cc40abd16b858e8aa3b0e6f452789b4b28b Mon Sep 17 00:00:00 2001 From: Artie Khovanov Date: Fri, 3 Jan 2025 07:21:52 +0000 Subject: [PATCH 103/189] feat(Algebra/Group/Even): "Advanced" lemmas about even elements. (#20272) Add construction of subgroup of even elements / squares. Add result that squares (`IsSquare`) are non-negative. These results cannot be added to `Mathlib.Algebra.Group.Even` directly because of import restrictions. This PR is split off from #16094 --- Mathlib.lean | 1 + Mathlib/Algebra/Group/Subgroup/Even.lean | 85 ++++++++++++++++++++++++ Mathlib/Algebra/Order/Ring/Basic.lean | 6 ++ Mathlib/Tactic/ToAdditive/Frontend.lean | 1 + 4 files changed, 93 insertions(+) create mode 100644 Mathlib/Algebra/Group/Subgroup/Even.lean diff --git a/Mathlib.lean b/Mathlib.lean index af4002f52b026..17d06a188c046 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -319,6 +319,7 @@ import Mathlib.Algebra.Group.Semiconj.Units import Mathlib.Algebra.Group.Subgroup.Actions import Mathlib.Algebra.Group.Subgroup.Basic import Mathlib.Algebra.Group.Subgroup.Defs +import Mathlib.Algebra.Group.Subgroup.Even import Mathlib.Algebra.Group.Subgroup.Finite import Mathlib.Algebra.Group.Subgroup.Finsupp import Mathlib.Algebra.Group.Subgroup.Ker diff --git a/Mathlib/Algebra/Group/Subgroup/Even.lean b/Mathlib/Algebra/Group/Subgroup/Even.lean new file mode 100644 index 0000000000000..b4735654ebe05 --- /dev/null +++ b/Mathlib/Algebra/Group/Subgroup/Even.lean @@ -0,0 +1,85 @@ +/- +Copyright (c) 2024 Artie Khovanov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Artie Khovanov +-/ +import Mathlib.Algebra.Group.Even +import Mathlib.Algebra.Group.Subgroup.Defs + +/-! +# Squares and even elements + +This file defines the subgroup of squares / even elements in an abelian group. +-/ + +namespace Subsemigroup +variable {S : Type*} [CommSemigroup S] {a : S} + +variable (S) in +/-- +In a commutative semigroup `S`, `Subsemigroup.squareIn S` is the subsemigroup of squares in `S`. +-/ +@[to_additive +"In a commutative additive semigroup `S`, the type `AddSubsemigroup.evenIn S` +is the subsemigroup of even elements of `S`."] +def squareIn : Subsemigroup S where + carrier := {s : S | IsSquare s} + mul_mem' := IsSquare.mul + +@[to_additive (attr := simp)] +theorem mem_squareIn : a ∈ squareIn S ↔ IsSquare a := Iff.rfl + +@[to_additive (attr := simp, norm_cast)] +theorem coe_squareIn : squareIn S = {s : S | IsSquare s} := rfl + +end Subsemigroup + +namespace Submonoid +variable {M : Type*} [CommMonoid M] {a : M} + +variable (M) in +/-- +In a commutative monoid `M`, `Submonoid.squareIn M` is the submonoid of squares in `M`. +-/ +@[to_additive +"In a commutative additive monoid `M`, the type `AddSubmonoid.evenIn M` +is the submonoid of even elements of `M`."] +def squareIn : Submonoid M where + __ := Subsemigroup.squareIn M + one_mem' := IsSquare.one + +@[to_additive (attr := simp)] +theorem squareIn_toSubsemigroup : (squareIn M).toSubsemigroup = .squareIn M := rfl + +@[to_additive (attr := simp)] +theorem mem_squareIn : a ∈ squareIn M ↔ IsSquare a := Iff.rfl + +@[to_additive (attr := simp, norm_cast)] +theorem coe_squareIn : squareIn M = {s : M | IsSquare s} := rfl + +end Submonoid + +namespace Subgroup +variable {G : Type*} [CommGroup G] {a : G} + +variable (G) in +/-- +In an abelian group `G`, `Subgroup.squareIn G` is the subgroup of squares in `G`. +-/ +@[to_additive +"In an abelian additive group `G`, the type `AddSubgroup.evenIn G` is +the subgroup of even elements in `G`."] +def squareIn : Subgroup G where + __ := Submonoid.squareIn G + inv_mem' := IsSquare.inv + +@[to_additive (attr := simp)] +theorem squareIn_toSubmonoid : (squareIn G).toSubmonoid = .squareIn G := rfl + +@[to_additive (attr := simp)] +theorem mem_squareIn : a ∈ squareIn G ↔ IsSquare a := Iff.rfl + +@[to_additive (attr := simp, norm_cast)] +theorem coe_squareIn : squareIn G = {s : G | IsSquare s} := rfl + +end Subgroup diff --git a/Mathlib/Algebra/Order/Ring/Basic.lean b/Mathlib/Algebra/Order/Ring/Basic.lean index 6f6a50a96337c..5085a583c3ee7 100644 --- a/Mathlib/Algebra/Order/Ring/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Basic.lean @@ -20,6 +20,12 @@ open Function Int variable {α M R : Type*} +theorem IsSquare.nonneg [Semiring R] [LinearOrder R] [IsRightCancelAdd R] + [ZeroLEOneClass R] [ExistsAddOfLE R] [PosMulMono R] [AddLeftStrictMono R] + {x : R} (h : IsSquare x) : 0 ≤ x := by + rcases h with ⟨y, rfl⟩ + exact mul_self_nonneg y + namespace MonoidHom variable [Ring R] [Monoid M] [LinearOrder M] [MulLeftMono M] (f : R →* M) diff --git a/Mathlib/Tactic/ToAdditive/Frontend.lean b/Mathlib/Tactic/ToAdditive/Frontend.lean index 6e1942fba38c6..108924443e60e 100644 --- a/Mathlib/Tactic/ToAdditive/Frontend.lean +++ b/Mathlib/Tactic/ToAdditive/Frontend.lean @@ -1027,6 +1027,7 @@ def fixAbbreviation : List String → List String | "add" :: "_" :: "indicator" :: s => "indicator" :: fixAbbreviation s | "is" :: "Square" :: s => "even" :: fixAbbreviation s | "Is" :: "Square" :: s => "Even" :: fixAbbreviation s + | "square" :: "In" :: s => "evenIn" :: fixAbbreviation s -- "Regular" is well-used in mathlib with various meanings (e.g. in -- measure theory) and a direct translation -- "regular" --> ["add", "Regular"] in `nameDict` above seems error-prone. From 9082fa7b9f7bcad5b4b6052a629eb9347580073a Mon Sep 17 00:00:00 2001 From: damiano Date: Fri, 3 Jan 2025 09:11:20 +0000 Subject: [PATCH 104/189] chore(Linter/UnusedTactic): do not exclude `change` from linted tactics (#20412) An earlier implementation of the `unusedTactic` linter tested for definitional equality. For this reason, the tactic `change` was excluded from the list of tactics that the linter inspected. This PR removes the custom exception on `change`, since the linter actually checks changes in the metavariables themselves, rather than definitional equality of the goals. As a consequence, some `change`s present in mathlib can be removed: in all cases, I verified that the infoview shows no difference before and after `change`. [Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Persistent.20.23allow_unused_tactic) --- Mathlib/Algebra/EuclideanDomain/Defs.lean | 7 ++----- Mathlib/Algebra/Polynomial/RingDivision.lean | 1 - Mathlib/Algebra/Ring/WithZero.lean | 3 +-- Mathlib/CategoryTheory/Monad/Comonadicity.lean | 4 +--- Mathlib/CategoryTheory/Monad/Monadicity.lean | 4 +--- Mathlib/CategoryTheory/Sites/Sieves.lean | 1 - .../SimpleGraph/Regularity/Chunk.lean | 1 - Mathlib/GroupTheory/DoubleCoset.lean | 2 -- Mathlib/NumberTheory/Padics/PadicNumbers.lean | 3 +-- Mathlib/Tactic/Linter/UnusedTactic.lean | 1 - Mathlib/Topology/Partial.lean | 1 - MathlibTest/UnusedTactic.lean | 17 +++++++++++++++++ MathlibTest/apply_fun.lean | 1 + 13 files changed, 24 insertions(+), 22 deletions(-) diff --git a/Mathlib/Algebra/EuclideanDomain/Defs.lean b/Mathlib/Algebra/EuclideanDomain/Defs.lean index 429e46451dffe..becb2fd79dcfe 100644 --- a/Mathlib/Algebra/EuclideanDomain/Defs.lean +++ b/Mathlib/Algebra/EuclideanDomain/Defs.lean @@ -159,11 +159,8 @@ section theorem GCD.induction {P : R → R → Prop} (a b : R) (H0 : ∀ x, P 0 x) (H1 : ∀ a b, a ≠ 0 → P (b % a) a → P a b) : P a b := by classical - exact if a0 : a = 0 then by - -- Porting note: required for hygiene, the equation compiler introduces a dummy variable `x` - -- See https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/unnecessarily.20tombstoned.20argument/near/314573315 - change P a b - exact a0.symm ▸ H0 b + exact if a0 : a = 0 then + a0.symm ▸ H0 b else have _ := mod_lt b a0 H1 _ _ a0 (GCD.induction (b % a) a H0 H1) diff --git a/Mathlib/Algebra/Polynomial/RingDivision.lean b/Mathlib/Algebra/Polynomial/RingDivision.lean index ad3c8e6a0e543..cf24ecbffc9f9 100644 --- a/Mathlib/Algebra/Polynomial/RingDivision.lean +++ b/Mathlib/Algebra/Polynomial/RingDivision.lean @@ -267,7 +267,6 @@ theorem exists_multiset_roots [DecidableEq R] : rw [← degree_add_divByMonic (monic_X_sub_C x) hdeg, degree_X_sub_C, add_comm] exact add_le_add (le_refl (1 : WithBot ℕ)) htd, by - change ∀ (a : R), count a (x ::ₘ t) = rootMultiplicity a p intro a conv_rhs => rw [← mul_divByMonic_eq_iff_isRoot.mpr hx] rw [rootMultiplicity_mul (mul_ne_zero (X_sub_C_ne_zero x) hdiv0), diff --git a/Mathlib/Algebra/Ring/WithZero.lean b/Mathlib/Algebra/Ring/WithZero.lean index 6c68ee84556bc..4acfa8d6aa5c3 100644 --- a/Mathlib/Algebra/Ring/WithZero.lean +++ b/Mathlib/Algebra/Ring/WithZero.lean @@ -24,8 +24,7 @@ instance instRightDistribClass [Mul α] [Add α] [RightDistribClass α] : RightDistribClass (WithZero α) where right_distrib a b c := by cases' c with c - · change (a + b) * 0 = a * 0 + b * 0 - simp + · simp cases' a with a <;> cases' b with b <;> try rfl exact congr_arg some (right_distrib _ _ _) diff --git a/Mathlib/CategoryTheory/Monad/Comonadicity.lean b/Mathlib/CategoryTheory/Monad/Comonadicity.lean index d72e8b6e22780..c7a1009392292 100644 --- a/Mathlib/CategoryTheory/Monad/Comonadicity.lean +++ b/Mathlib/CategoryTheory/Monad/Comonadicity.lean @@ -146,9 +146,7 @@ def counitFork (A : adj.toComonad.Coalgebra) [HasEqualizer (G.map A.a) (adj.unit.app (G.obj A.A))] : Fork (F.map (G.map A.a)) (F.map (adj.unit.app (G.obj A.A))) := Fork.ofι (F.map (equalizer.ι (G.map A.a) (adj.unit.app (G.obj A.A)))) - (by - change _ = F.map _ ≫ _ - rw [← F.map_comp, equalizer.condition, F.map_comp]) + (by rw [← F.map_comp, equalizer.condition, F.map_comp]) @[simp] theorem unitFork_ι (A : adj.toComonad.Coalgebra) diff --git a/Mathlib/CategoryTheory/Monad/Monadicity.lean b/Mathlib/CategoryTheory/Monad/Monadicity.lean index bd9b396629363..591201fa31725 100644 --- a/Mathlib/CategoryTheory/Monad/Monadicity.lean +++ b/Mathlib/CategoryTheory/Monad/Monadicity.lean @@ -160,9 +160,7 @@ def unitCofork (A : adj.toMonad.Algebra) [HasCoequalizer (F.map A.a) (adj.counit.app (F.obj A.A))] : Cofork (G.map (F.map A.a)) (G.map (adj.counit.app (F.obj A.A))) := Cofork.ofπ (G.map (coequalizer.π (F.map A.a) (adj.counit.app (F.obj A.A)))) - (by - change _ = G.map _ ≫ _ - rw [← G.map_comp, coequalizer.condition, G.map_comp]) + (by rw [← G.map_comp, coequalizer.condition, G.map_comp]) @[simp] theorem unitCofork_π (A : adj.toMonad.Algebra) diff --git a/Mathlib/CategoryTheory/Sites/Sieves.lean b/Mathlib/CategoryTheory/Sites/Sieves.lean index f2446cb5f188c..23b36b70f4d19 100644 --- a/Mathlib/CategoryTheory/Sites/Sieves.lean +++ b/Mathlib/CategoryTheory/Sites/Sieves.lean @@ -626,7 +626,6 @@ theorem pullbackArrows_comm [HasPullbacks C] {X Y : C} (f : Y ⟶ X) (R : Presie constructor · rintro ⟨_, h, k, hk, rfl⟩ cases' hk with W g hg - change (Sieve.generate R).pullback f (h ≫ pullback.snd g f) rw [Sieve.pullback_apply, assoc, ← pullback.condition, ← assoc] exact Sieve.downward_closed _ (by exact Sieve.le_generate R W hg) (h ≫ pullback.fst g f) · rintro ⟨W, h, k, hk, comm⟩ diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean index 7759eae0b284d..067be4ae05806 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean @@ -428,7 +428,6 @@ private theorem edgeDensity_star_not_uniform [Nonempty α] set q : ℝ := (∑ ab ∈ (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, (G.edgeDensity ab.1 ab.2 : ℝ)) / (↑4 ^ #P.parts * ↑4 ^ #P.parts) - change _ ≤ |p - q| set r : ℝ := ↑(G.edgeDensity ((star hP G ε hU V).biUnion id) ((star hP G ε hV U).biUnion id)) set s : ℝ := ↑(G.edgeDensity (G.nonuniformWitness ε U V) (G.nonuniformWitness ε V U)) set t : ℝ := ↑(G.edgeDensity U V) diff --git a/Mathlib/GroupTheory/DoubleCoset.lean b/Mathlib/GroupTheory/DoubleCoset.lean index 5efccbbe6d6e7..22b908c845a49 100644 --- a/Mathlib/GroupTheory/DoubleCoset.lean +++ b/Mathlib/GroupTheory/DoubleCoset.lean @@ -83,7 +83,6 @@ theorem bot_rel_eq_leftRel (H : Subgroup G) : rw [rel_iff, QuotientGroup.leftRel_apply] constructor · rintro ⟨a, rfl : a = 1, b, hb, rfl⟩ - change a⁻¹ * (1 * a * b) ∈ H rwa [one_mul, inv_mul_cancel_left] · rintro (h : a⁻¹ * b ∈ H) exact ⟨1, rfl, a⁻¹ * b, h, by rw [one_mul, mul_inv_cancel_left]⟩ @@ -94,7 +93,6 @@ theorem rel_bot_eq_right_group_rel (H : Subgroup G) : rw [rel_iff, QuotientGroup.rightRel_apply] constructor · rintro ⟨b, hb, a, rfl : a = 1, rfl⟩ - change b * a * 1 * a⁻¹ ∈ H rwa [mul_one, mul_inv_cancel_right] · rintro (h : b * a⁻¹ ∈ H) exact ⟨b * a⁻¹, h, 1, rfl, by rw [mul_one, inv_mul_cancel_right]⟩ diff --git a/Mathlib/NumberTheory/Padics/PadicNumbers.lean b/Mathlib/NumberTheory/Padics/PadicNumbers.lean index e961a395e031e..a2b72fdf7bbad 100644 --- a/Mathlib/NumberTheory/Padics/PadicNumbers.lean +++ b/Mathlib/NumberTheory/Padics/PadicNumbers.lean @@ -952,8 +952,7 @@ theorem norm_eq_zpow_neg_valuation {x : ℚ_[p]} : x ≠ 0 → ‖x‖ = (p : refine Quotient.inductionOn' x fun f hf => ?_ change (PadicSeq.norm _ : ℝ) = (p : ℝ) ^ (-PadicSeq.valuation _) rw [PadicSeq.norm_eq_zpow_neg_valuation] - · change ↑((p : ℚ) ^ (-PadicSeq.valuation f)) = (p : ℝ) ^ (-PadicSeq.valuation f) - rw [Rat.cast_zpow, Rat.cast_natCast] + · rw [Rat.cast_zpow, Rat.cast_natCast] · apply CauSeq.not_limZero_of_not_congr_zero -- Porting note: was `contrapose! hf` intro hf' diff --git a/Mathlib/Tactic/Linter/UnusedTactic.lean b/Mathlib/Tactic/Linter/UnusedTactic.lean index 01c6ba0915ce9..807699e1024c3 100644 --- a/Mathlib/Tactic/Linter/UnusedTactic.lean +++ b/Mathlib/Tactic/Linter/UnusedTactic.lean @@ -93,7 +93,6 @@ initialize allowedRef : IO.Ref (Std.HashSet SyntaxNodeKind) ← `Mathlib.Tactic.Propose.propose', `Lean.Parser.Tactic.traceState, `Mathlib.Tactic.tacticMatch_target_, - ``Lean.Parser.Tactic.change, `change?, `«tactic#adaptation_note_», `tacticSleep_heartbeats_, diff --git a/Mathlib/Topology/Partial.lean b/Mathlib/Topology/Partial.lean index 5c1301d71d438..099830f1bc2a4 100644 --- a/Mathlib/Topology/Partial.lean +++ b/Mathlib/Topology/Partial.lean @@ -62,7 +62,6 @@ theorem pcontinuous_iff' {f : X →. Y} : rintro x ⟨y, ys, fxy⟩ t rw [mem_principal] intro (h : f.preimage s ⊆ t) - change t ∈ 𝓝 x apply mem_of_superset _ h have h' : ∀ s ∈ 𝓝 y, f.preimage s ∈ 𝓝 x := by intro s hs diff --git a/MathlibTest/UnusedTactic.lean b/MathlibTest/UnusedTactic.lean index e81739bb01e4d..c3614b7fc021d 100644 --- a/MathlibTest/UnusedTactic.lean +++ b/MathlibTest/UnusedTactic.lean @@ -1,6 +1,23 @@ import Mathlib.Tactic.Linter.UnusedTactic import Mathlib.Tactic.AdaptationNote +example (h : 0 + 1 = 0) : False := by + change 1 = 0 at h + simp at h + +example : 0 + 1 = 1 := by + change 1 = 1 + rfl + +/-- +warning: 'change 1 = 1' tactic does nothing +note: this linter can be disabled with `set_option linter.unusedTactic false` +-/ +#guard_msgs in +example : 1 = 1 := by + change 1 = 1 + rfl + def why2 : True → True := (by refine ·) example : True := by diff --git a/MathlibTest/apply_fun.lean b/MathlibTest/apply_fun.lean index bb23b864e3ef2..7c9ff9b2b5b0c 100644 --- a/MathlibTest/apply_fun.lean +++ b/MathlibTest/apply_fun.lean @@ -234,6 +234,7 @@ example (α β : Type u) [Fintype α] [Fintype β] (h : α = β) : True := by trivial -- Check that metavariables in the goal do not prevent apply_fun from detecting the relation +set_option linter.unusedTactic false in example (f : α ≃ β) (x y : α) (h : f x = f y) : x = y := by change _ -- now the goal is a metavariable From 3997f6f56c915bc9ea794c3b72bfe0cf3966a97f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 3 Jan 2025 11:02:32 +0000 Subject: [PATCH 105/189] =?UTF-8?q?feat:=20`f=20''=20s=E1=B6=9C=20=3D=20ra?= =?UTF-8?q?nge=20f=20\=20f=20''=20s`=20when=20`f`=20is=20injective=20(#190?= =?UTF-8?q?30)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From GrowthInGroups (LeanCamCombi) --- Mathlib/Data/Set/Image.lean | 11 +++++++---- 1 file changed, 7 insertions(+), 4 deletions(-) diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 2df34e18001a6..a98696d44ae14 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -610,6 +610,13 @@ theorem image_univ {f : α → β} : f '' univ = range f := by ext simp [image, range] +lemma image_compl_eq_range_diff_image {f : α → β} (hf : Injective f) (s : Set α) : + f '' sᶜ = range f \ f '' s := by rw [← image_univ, ← image_diff hf, compl_eq_univ_diff] + +/-- Alias of `Set.image_compl_eq_range_sdiff_image`. -/ +lemma range_diff_image {f : α → β} (hf : Injective f) (s : Set α) : range f \ f '' s = f '' sᶜ := by + rw [image_compl_eq_range_diff_image hf] + @[simp] theorem preimage_eq_univ_iff {f : α → β} {s} : f ⁻¹' s = univ ↔ range f ⊆ s := by rw [← univ_subset_iff, ← image_subset_iff, image_univ] @@ -942,10 +949,6 @@ theorem range_unique [h : Unique ι] : range f = {f default} := by theorem range_diff_image_subset (f : α → β) (s : Set α) : range f \ f '' s ⊆ f '' sᶜ := fun _ ⟨⟨x, h₁⟩, h₂⟩ => ⟨x, fun h => h₂ ⟨x, h, h₁⟩, h₁⟩ -theorem range_diff_image {f : α → β} (H : Injective f) (s : Set α) : range f \ f '' s = f '' sᶜ := - (Subset.antisymm (range_diff_image_subset f s)) fun _ ⟨_, hx, hy⟩ => - hy ▸ ⟨mem_range_self _, fun ⟨_, hx', Eq⟩ => hx <| H Eq ▸ hx'⟩ - @[simp] theorem range_inclusion (h : s ⊆ t) : range (inclusion h) = { x : t | (x : α) ∈ s } := by ext ⟨x, hx⟩ From 08ad1629fff207dfed86280ee5efbb9c7d8e850e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 3 Jan 2025 11:02:33 +0000 Subject: [PATCH 106/189] feat(Probability): proper kernels (#19434) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define proper kernels, which are self-kernels which can meaningfully be restricted to an event. From GibbsMeasure Co-authored-by: Kalle Kytölä Co-authored-by: Kin Yau James Wong --- Mathlib.lean | 1 + Mathlib/Probability/Kernel/Proper.lean | 139 +++++++++++++++++++++++++ 2 files changed, 140 insertions(+) create mode 100644 Mathlib/Probability/Kernel/Proper.lean diff --git a/Mathlib.lean b/Mathlib.lean index 17d06a188c046..fb780d7561c21 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4269,6 +4269,7 @@ import Mathlib.Probability.Kernel.Disintegration.Unique import Mathlib.Probability.Kernel.Integral import Mathlib.Probability.Kernel.Invariance import Mathlib.Probability.Kernel.MeasurableIntegral +import Mathlib.Probability.Kernel.Proper import Mathlib.Probability.Kernel.RadonNikodym import Mathlib.Probability.Kernel.WithDensity import Mathlib.Probability.Martingale.Basic diff --git a/Mathlib/Probability/Kernel/Proper.lean b/Mathlib/Probability/Kernel/Proper.lean new file mode 100644 index 0000000000000..b7cb185ae7528 --- /dev/null +++ b/Mathlib/Probability/Kernel/Proper.lean @@ -0,0 +1,139 @@ +/- +Copyright (c) 2024 Yaël Dillies, Kalle Kytölä, Kin Yau James Wong. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Kalle Kytölä, Kin Yau James Wong +-/ +import Mathlib.Probability.Kernel.Basic + +/-! +# Proper kernels + +This file defines properness of measure kernels. + +For two σ-algebras `𝓑 ≤ 𝓧`, a `𝓑, 𝓧`-kernel `π : X → Measure X` is proper if +`∫ x, g x * f x ∂(π x₀) = g x₀ * ∫ x, f x ∂(π x₀)` for all `x₀ : X`, `𝓧`-measurable function `f` +and `𝓑`-measurable function `g`. + +By the standard machine, this is equivalent to having that, for all `B ∈ 𝓑`, `π` restricted to `B` +is the same as `π` times the indicator of `B`. + +This should be thought of as the condition under which one can meaningfully restrict a kernel to an +event. + +## TODO + +Prove the `integral` versions of the `lintegral` lemmas below +-/ + +open MeasureTheory ENNReal NNReal Set +open scoped ProbabilityTheory + +namespace ProbabilityTheory.Kernel +variable {X : Type*} {𝓑 𝓧 : MeasurableSpace X} {π : Kernel[𝓑, 𝓧] X X} {A B : Set X} + {f g : X → ℝ≥0∞} {x₀ : X} + +/-- For two σ-algebras `𝓑 ≤ 𝓧` on a space `X`, a `𝓑, 𝓧`-kernel `π : X → Measure X` is proper if +`∫ x, g x * f x ∂(π x₀) = g x₀ * ∫ x, f x ∂(π x₀)` for all `x₀ : X`, `𝓧`-measurable function `f` +and `𝓑`-measurable function `g`. + +By the standard machine, this is equivalent to having that, for all `B ∈ 𝓑`, `π` restricted to `B` +is the same as `π` times the indicator of `B`. + +To avoid assuming `𝓑 ≤ 𝓧` in the definition, we replace `𝓑` by `𝓑 ⊓ 𝓧` in the restriction. -/ +structure IsProper (π : Kernel[𝓑, 𝓧] X X) : Prop where + restrict_eq_indicator_smul' : + ∀ ⦃B : Set X⦄ (hB : MeasurableSet[𝓑 ⊓ 𝓧] B) (x : X), + π.restrict (inf_le_right (b := 𝓧) _ hB) x = B.indicator (fun _ ↦ (1 : ℝ≥0∞)) x • π x + +lemma isProper_iff_restrict_eq_indicator_smul (h𝓑𝓧 : 𝓑 ≤ 𝓧) : + IsProper π ↔ ∀ ⦃B : Set X⦄ (hB : MeasurableSet[𝓑] B) (x : X), + π.restrict (h𝓑𝓧 _ hB) x = B.indicator (fun _ ↦ (1 : ℝ≥0∞)) x • π x := by + refine ⟨fun ⟨h⟩ ↦ ?_, fun h ↦ ⟨?_⟩⟩ <;> simpa only [inf_eq_left.2 h𝓑𝓧] using h + +lemma isProper_iff_inter_eq_indicator_mul (h𝓑𝓧 : 𝓑 ≤ 𝓧) : + IsProper π ↔ + ∀ ⦃A : Set X⦄ (_hA : MeasurableSet[𝓧] A) ⦃B : Set X⦄ (_hB : MeasurableSet[𝓑] B) (x : X), + π x (A ∩ B) = B.indicator 1 x * π x A := by + calc + _ ↔ ∀ ⦃A : Set X⦄ (_hA : MeasurableSet[𝓧] A) ⦃B : Set X⦄ (hB : MeasurableSet[𝓑] B) (x : X), + π.restrict (h𝓑𝓧 _ hB) x A = B.indicator 1 x * π x A := by + simp [isProper_iff_restrict_eq_indicator_smul h𝓑𝓧, Measure.ext_iff]; aesop + _ ↔ _ := by congr! 5 with A hA B hB x; rw [restrict_apply, Measure.restrict_apply hA] + +alias ⟨IsProper.restrict_eq_indicator_smul, IsProper.of_restrict_eq_indicator_smul⟩ := + isProper_iff_restrict_eq_indicator_smul + +alias ⟨IsProper.inter_eq_indicator_mul, IsProper.of_inter_eq_indicator_mul⟩ := + isProper_iff_inter_eq_indicator_mul + +lemma IsProper.setLIntegral_eq_bind (hπ : IsProper π) (h𝓑𝓧 : 𝓑 ≤ 𝓧) {μ : Measure[𝓧] X} + (hA : MeasurableSet[𝓧] A) (hB : MeasurableSet[𝓑] B) : + ∫⁻ a in B, π a A ∂μ = μ.bind π (A ∩ B) := by + rw [Measure.bind_apply (by measurability) (π.measurable.mono h𝓑𝓧 le_rfl)] + simp only [hπ.inter_eq_indicator_mul h𝓑𝓧 hA hB, ← indicator_mul_const, Pi.one_apply, one_mul] + rw [← lintegral_indicator (h𝓑𝓧 _ hB)] + rfl + +/-- Auxiliary lemma for `IsProper.lintegral_mul` and +`IsProper.setLIntegral_eq_indicator_mul_lintegral`. -/ +private lemma IsProper.lintegral_indicator_mul_indicator (hπ : IsProper π) (h𝓑𝓧 : 𝓑 ≤ 𝓧) + (hA : MeasurableSet[𝓧] A) (hB : MeasurableSet[𝓑] B) : + ∫⁻ x, B.indicator 1 x * A.indicator 1 x ∂(π x₀) = + B.indicator 1 x₀ * ∫⁻ x, A.indicator 1 x ∂(π x₀) := by + simp_rw [← inter_indicator_mul] + rw [lintegral_indicator ((h𝓑𝓧 _ hB).inter hA), lintegral_indicator hA] + simp only [MeasureTheory.lintegral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter, + Pi.one_apply, one_mul] + rw [← hπ.inter_eq_indicator_mul h𝓑𝓧 hA hB, inter_comm] + +set_option linter.style.multiGoal false in -- false positive +/-- Auxiliary lemma for `IsProper.lintegral_mul` and +`IsProper.setLIntegral_eq_indicator_mul_lintegral`. -/ +private lemma IsProper.lintegral_indicator_mul (hπ : IsProper π) (h𝓑𝓧 : 𝓑 ≤ 𝓧) + (hf : Measurable[𝓧] f) (hB : MeasurableSet[𝓑] B) : + ∫⁻ x, B.indicator 1 x * f x ∂(π x₀) = B.indicator 1 x₀ * ∫⁻ x, f x ∂(π x₀) := by + refine hf.ennreal_induction ?_ ?_ ?_ + · rintro c A hA + simp_rw [← smul_indicator_one_apply, mul_smul_comm, smul_eq_mul] + rw [lintegral_const_mul, lintegral_const_mul, hπ.lintegral_indicator_mul_indicator h𝓑𝓧 hA hB, + mul_left_comm] <;> measurability + · rintro f₁ f₂ - _ _ hf₁ hf₂ + simp only [Pi.add_apply, mul_add] + rw [lintegral_add_right, lintegral_add_right, hf₁, hf₂, mul_add] <;> measurability + · rintro f' hf'_meas hf'_mono hf' + simp_rw [ENNReal.mul_iSup] + rw [lintegral_iSup (by measurability), lintegral_iSup hf'_meas hf'_mono, ENNReal.mul_iSup] + simp_rw [hf'] + · exact hf'_mono.const_mul (zero_le _) + +lemma IsProper.setLIntegral_eq_indicator_mul_lintegral (hπ : IsProper π) (h𝓑𝓧 : 𝓑 ≤ 𝓧) + (hf : Measurable[𝓧] f) (hB : MeasurableSet[𝓑] B) (x₀ : X) : + ∫⁻ x in B, f x ∂(π x₀) = B.indicator 1 x₀ * ∫⁻ x, f x ∂(π x₀) := by + simp [← hπ.lintegral_indicator_mul h𝓑𝓧 hf hB, ← indicator_mul_left, + lintegral_indicator (h𝓑𝓧 _ hB)] + +lemma IsProper.setLIntegral_inter_eq_indicator_mul_setLIntegral (hπ : IsProper π) (h𝓑𝓧 : 𝓑 ≤ 𝓧) + (hf : Measurable[𝓧] f) (hA : MeasurableSet[𝓧] A) (hB : MeasurableSet[𝓑] B) (x₀ : X) : + ∫⁻ x in A ∩ B, f x ∂(π x₀) = B.indicator 1 x₀ * ∫⁻ x in A, f x ∂(π x₀) := by + rw [← lintegral_indicator hA, ← hπ.setLIntegral_eq_indicator_mul_lintegral h𝓑𝓧 _ hB, + setLIntegral_indicator] <;> measurability + +lemma IsProper.lintegral_mul (hπ : IsProper π) (h𝓑𝓧 : 𝓑 ≤ 𝓧) (hf : Measurable[𝓧] f) + (hg : Measurable[𝓑] g) (x₀ : X) : + ∫⁻ x, g x * f x ∂(π x₀) = g x₀ * ∫⁻ x, f x ∂(π x₀) := by + refine hg.ennreal_induction ?_ ?_ ?_ + · rintro c A hA + simp_rw [← smul_indicator_one_apply, smul_mul_assoc, smul_eq_mul] + rw [lintegral_const_mul, hπ.lintegral_indicator_mul h𝓑𝓧 hf hA] + · measurability + · rintro g₁ g₂ - _ hg₂_meas hg₁ hg₂ + simp only [Pi.add_apply, mul_add, add_mul] + rw [lintegral_add_right, hg₁, hg₂] + · exact (hg₂_meas.mono h𝓑𝓧 le_rfl).mul hf + · rintro g' hg'_meas hg'_mono hg' + simp_rw [ENNReal.iSup_mul] + rw [lintegral_iSup (fun n ↦ ((hg'_meas _).mono h𝓑𝓧 le_rfl).mul hf) + (hg'_mono.mul_const (zero_le _))] + simp_rw [hg'] + +end ProbabilityTheory.Kernel From ef4bb8a4e81484f8edd70d2f51e6435ce76b233e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 3 Jan 2025 11:02:35 +0000 Subject: [PATCH 107/189] chore: turn left and right transversals from sets to types (#20041) The sets both duplicate `Subgroup.IsComplement`. --- Mathlib/GroupTheory/Complement.lean | 325 +++++++++++++++++------ Mathlib/GroupTheory/CosetCover.lean | 14 +- Mathlib/GroupTheory/HNNExtension.lean | 2 +- Mathlib/GroupTheory/PushoutI.lean | 2 +- Mathlib/GroupTheory/Schreier.lean | 46 ++-- Mathlib/GroupTheory/SchurZassenhaus.lean | 2 +- Mathlib/GroupTheory/Transfer.lean | 26 +- 7 files changed, 283 insertions(+), 134 deletions(-) diff --git a/Mathlib/GroupTheory/Complement.lean b/Mathlib/GroupTheory/Complement.lean index 9bad485b2ce2a..4425a82f957a6 100644 --- a/Mathlib/GroupTheory/Complement.lean +++ b/Mathlib/GroupTheory/Complement.lean @@ -12,21 +12,19 @@ In this file we define the complement of a subgroup. ## Main definitions -- `IsComplement S T` where `S` and `T` are subsets of `G` states that every `g : G` can be +- `Subgroup.IsComplement S T` where `S` and `T` are subsets of `G` states that every `g : G` can be written uniquely as a product `s * t` for `s ∈ S`, `t ∈ T`. -- `leftTransversals T` where `T` is a subset of `G` is the set of all left-complements of `T`, - i.e. the set of all `S : Set G` that contain exactly one element of each left coset of `T`. -- `rightTransversals S` where `S` is a subset of `G` is the set of all right-complements of `S`, - i.e. the set of all `T : Set G` that contain exactly one element of each right coset of `S`. -- `transferTransversal H g` is a specific `leftTransversal` of `H` that is used in the - computation of the transfer homomorphism evaluated at an element `g : G`. +- `H.LeftTransversal` where `H` is a subgroup of `G` is the type of all left-complements of `H`, + i.e. the set of all `S : Set G` that contain exactly one element of each left coset of `H`. +- `H.RightTransversal` where `H` is a subgroup of `G` is the set of all right-complements of `H`, + i.e. the set of all `T : Set G` that contain exactly one element of each right coset of `H`. ## Main results - `isComplement'_of_coprime` : Subgroups of coprime order are complements. -/ -open Set +open Function Set open scoped Pointwise namespace Subgroup @@ -45,12 +43,14 @@ abbrev IsComplement' := IsComplement (H : Set G) (K : Set G) /-- The set of left-complements of `T : Set G` -/ -@[to_additive "The set of left-complements of `T : Set G`"] +@[to_additive (attr := deprecated IsComplement (since := "2024-12-18")) +"The set of left-complements of `T : Set G`"] def leftTransversals : Set (Set G) := { S : Set G | IsComplement S T } /-- The set of right-complements of `S : Set G` -/ -@[to_additive "The set of right-complements of `S : Set G`"] +@[to_additive (attr := deprecated IsComplement (since := "2024-12-18")) +"The set of right-complements of `S : Set G`"] def rightTransversals : Set (Set G) := { T : Set G | IsComplement S T } @@ -169,6 +169,16 @@ theorem isComplement'_top_right : IsComplement' H ⊤ ↔ H = ⊥ := isComplement_univ_right.trans coe_eq_singleton @[to_additive] +lemma isComplement_iff_existsUnique_inv_mul_mem : + IsComplement S T ↔ ∀ g, ∃! s : S, (s : G)⁻¹ * g ∈ T := by + convert isComplement_iff_existsUnique with g + constructor <;> rintro ⟨x, hx, hx'⟩ + · exact ⟨(x, ⟨_, hx⟩), by simp, by aesop⟩ + · exact ⟨x.1, by simp [← hx], fun y hy ↦ (Prod.ext_iff.1 <| by simpa using hx' (y, ⟨_, hy⟩)).1⟩ + +set_option linter.deprecated false in +@[to_additive + (attr := deprecated isComplement_iff_existsUnique_inv_mul_mem (since := "2024-12-18"))] theorem mem_leftTransversals_iff_existsUnique_inv_mul_mem : S ∈ leftTransversals T ↔ ∀ g : G, ∃! s : S, (s : G)⁻¹ * g ∈ T := by rw [leftTransversals, Set.mem_setOf_eq, isComplement_iff_existsUnique] @@ -183,6 +193,16 @@ theorem mem_leftTransversals_iff_existsUnique_inv_mul_mem : exact Prod.ext hf (Subtype.ext (eq_inv_mul_of_mul_eq (hf ▸ hy))) @[to_additive] +lemma isComplement_iff_existsUnique_mul_inv_mem : + IsComplement S T ↔ ∀ g, ∃! t : T, g * (t : G)⁻¹ ∈ S := by + convert isComplement_iff_existsUnique with g + constructor <;> rintro ⟨x, hx, hx'⟩ + · exact ⟨(⟨_, hx⟩, x), by simp, by aesop⟩ + · exact ⟨x.2, by simp [← hx], fun y hy ↦ (Prod.ext_iff.1 <| by simpa using hx' (⟨_, hy⟩, y)).2⟩ + +set_option linter.deprecated false in +@[to_additive + (attr := deprecated isComplement_iff_existsUnique_mul_inv_mem (since := "2024-12-18"))] theorem mem_rightTransversals_iff_existsUnique_mul_inv_mem : S ∈ rightTransversals T ↔ ∀ g : G, ∃! s : S, g * (s : G)⁻¹ ∈ T := by rw [rightTransversals, Set.mem_setOf_eq, isComplement_iff_existsUnique] @@ -197,6 +217,15 @@ theorem mem_rightTransversals_iff_existsUnique_mul_inv_mem : exact Prod.ext (Subtype.ext (eq_mul_inv_of_mul_eq (hf ▸ hy))) hf @[to_additive] +lemma isComplement_subgroup_right_iff_existsUnique_quotientGroupMk : + IsComplement S H ↔ ∀ q : G ⧸ H, ∃! s : S, QuotientGroup.mk s.1 = q := by + simp_rw [isComplement_iff_existsUnique_inv_mul_mem, SetLike.mem_coe, ← QuotientGroup.eq, + QuotientGroup.forall_mk] + +set_option linter.deprecated false in +@[to_additive + (attr := deprecated isComplement_subgroup_right_iff_existsUnique_quotientGroupMk + (since := "2024-12-18"))] 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 @@ -204,7 +233,18 @@ theorem mem_leftTransversals_iff_existsUnique_quotient_mk''_eq : QuotientGroup.eq] exact ⟨fun h q => Quotient.inductionOn' q h, fun h g => h (Quotient.mk'' g)⟩ +set_option linter.docPrime false in @[to_additive] +lemma isComplement_subgroup_left_iff_existsUnique_quotientMk'' : + IsComplement H T ↔ + ∀ q : Quotient (QuotientGroup.rightRel H), ∃! t : T, Quotient.mk'' t.1 = q := by + simp_rw [isComplement_iff_existsUnique_mul_inv_mem, SetLike.mem_coe, + ← QuotientGroup.rightRel_apply, ← Quotient.eq'', Quotient.forall] + +set_option linter.deprecated false in +@[to_additive + (attr := deprecated isComplement_subgroup_left_iff_existsUnique_quotientMk'' + (since := "2024-12-18"))] theorem mem_rightTransversals_iff_existsUnique_quotient_mk''_eq : S ∈ rightTransversals (H : Set G) ↔ ∀ q : Quotient (QuotientGroup.rightRel H), ∃! s : S, Quotient.mk'' s.1 = q := by @@ -213,6 +253,14 @@ theorem mem_rightTransversals_iff_existsUnique_quotient_mk''_eq : exact ⟨fun h q => Quotient.inductionOn' q h, fun h g => h (Quotient.mk'' g)⟩ @[to_additive] +lemma isComplement_subgroup_right_iff_bijective : + IsComplement S H ↔ Bijective (S.restrict (QuotientGroup.mk : G → G ⧸ H)) := + isComplement_subgroup_right_iff_existsUnique_quotientGroupMk.trans + (bijective_iff_existsUnique (S.restrict QuotientGroup.mk)).symm + +set_option linter.deprecated false in +@[to_additive + (attr := deprecated isComplement_subgroup_right_iff_bijective (since := "2024-12-18"))] theorem mem_leftTransversals_iff_bijective : S ∈ leftTransversals (H : Set G) ↔ Function.Bijective (S.restrict (Quotient.mk'' : G → Quotient (QuotientGroup.leftRel H))) := @@ -220,6 +268,15 @@ theorem mem_leftTransversals_iff_bijective : (Function.bijective_iff_existsUnique (S.restrict Quotient.mk'')).symm @[to_additive] +lemma isComplement_subgroup_left_iff_bijective : + IsComplement H T ↔ + Bijective (T.restrict (Quotient.mk'' : G → Quotient (QuotientGroup.rightRel H))) := + isComplement_subgroup_left_iff_existsUnique_quotientMk''.trans + (bijective_iff_existsUnique (T.restrict Quotient.mk'')).symm + +set_option linter.deprecated false in +@[to_additive + (attr := deprecated isComplement_subgroup_left_iff_bijective (since := "2024-12-18"))] theorem mem_rightTransversals_iff_bijective : S ∈ rightTransversals (H : Set G) ↔ Function.Bijective (S.restrict (Quotient.mk'' : G → Quotient (QuotientGroup.rightRel H))) := @@ -227,16 +284,36 @@ theorem mem_rightTransversals_iff_bijective : (Function.bijective_iff_existsUnique (S.restrict Quotient.mk'')).symm @[to_additive] +lemma IsComplement.card_left (h : IsComplement S H) : Nat.card S = H.index := + Nat.card_congr <| .ofBijective _ <| isComplement_subgroup_right_iff_bijective.mp h + +set_option linter.deprecated false in +@[to_additive (attr := deprecated IsComplement.card_left (since := "2024-12-18"))] theorem card_left_transversal (h : S ∈ leftTransversals (H : Set G)) : Nat.card S = H.index := Nat.card_congr <| Equiv.ofBijective _ <| mem_leftTransversals_iff_bijective.mp h @[to_additive] +lemma IsComplement.card_right (h : IsComplement H T) : Nat.card T = H.index := + Nat.card_congr <| (Equiv.ofBijective _ <| isComplement_subgroup_left_iff_bijective.mp h).trans <| + QuotientGroup.quotientRightRelEquivQuotientLeftRel H + +set_option linter.deprecated false in +@[to_additive (attr := deprecated IsComplement.card_right (since := "2024-12-18"))] theorem card_right_transversal (h : S ∈ rightTransversals (H : Set G)) : Nat.card S = H.index := Nat.card_congr <| (Equiv.ofBijective _ <| mem_rightTransversals_iff_bijective.mp h).trans <| QuotientGroup.quotientRightRelEquivQuotientLeftRel H @[to_additive] +lemma isComplement_range_left {f : G ⧸ H → G} (hf : ∀ q, ↑(f q) = q) : + IsComplement (range f) H := by + rw [isComplement_subgroup_right_iff_bijective] + refine ⟨?_, fun q ↦ ⟨⟨f q, q, rfl⟩, hf q⟩⟩ + rintro ⟨-, q₁, rfl⟩ ⟨-, q₂, rfl⟩ h + exact Subtype.ext <| congr_arg f <| ((hf q₁).symm.trans h).trans (hf q₂) + +set_option linter.deprecated false in +@[to_additive (attr := deprecated isComplement_range_left (since := "2024-12-18"))] theorem range_mem_leftTransversals {f : G ⧸ H → G} (hf : ∀ q, ↑(f q) = q) : Set.range f ∈ leftTransversals (H : Set G) := mem_leftTransversals_iff_bijective.mpr @@ -245,6 +322,15 @@ theorem range_mem_leftTransversals {f : G ⧸ H → G} (hf : ∀ q, ↑(f q) = q fun q => ⟨⟨f q, q, rfl⟩, hf q⟩⟩ @[to_additive] +lemma isComplement_range_right {f : Quotient (QuotientGroup.rightRel H) → G} + (hf : ∀ q, Quotient.mk'' (f q) = q) : IsComplement H (range f) := by + rw [isComplement_subgroup_left_iff_bijective] + refine ⟨?_, fun q ↦ ⟨⟨f q, q, rfl⟩, hf q⟩⟩ + rintro ⟨-, q₁, rfl⟩ ⟨-, q₂, rfl⟩ h + exact Subtype.ext <| congr_arg f <| ((hf q₁).symm.trans h).trans (hf q₂) + +set_option linter.deprecated false in +@[to_additive (attr := deprecated isComplement_range_right (since := "2024-12-18"))] theorem range_mem_rightTransversals {f : Quotient (QuotientGroup.rightRel H) → G} (hf : ∀ q, Quotient.mk'' (f q) = q) : Set.range f ∈ rightTransversals (H : Set G) := mem_rightTransversals_iff_bijective.mpr @@ -253,6 +339,17 @@ theorem range_mem_rightTransversals {f : Quotient (QuotientGroup.rightRel H) → fun q => ⟨⟨f q, q, rfl⟩, hf q⟩⟩ @[to_additive] +lemma exists_isComplement_left (H : Subgroup G) (g : G) : ∃ S, IsComplement S H ∧ g ∈ S := by + classical + refine ⟨Set.range (Function.update Quotient.out _ g), isComplement_range_left fun q ↦ ?_, + QuotientGroup.mk g, Function.update_self (Quotient.mk'' g) g Quotient.out⟩ + by_cases hq : q = Quotient.mk'' g + · exact hq.symm ▸ congr_arg _ (Function.update_self (Quotient.mk'' g) g Quotient.out) + · refine Function.update_of_ne ?_ g Quotient.out ▸ q.out_eq' + exact hq + +set_option linter.deprecated false in +@[to_additive (attr := deprecated exists_isComplement_left (since := "2024-12-18"))] lemma exists_left_transversal (H : Subgroup G) (g : G) : ∃ S ∈ leftTransversals (H : Set G), g ∈ S := by classical @@ -265,6 +362,18 @@ lemma exists_left_transversal (H : Subgroup G) (g : G) : exact hq @[to_additive] +lemma exists_isComplement_right (H : Subgroup G) (g : G) : + ∃ T, IsComplement H T ∧ g ∈ T := by + classical + refine ⟨Set.range (Function.update Quotient.out _ g), isComplement_range_right fun q ↦ ?_, + Quotient.mk'' g, Function.update_self (Quotient.mk'' g) g Quotient.out⟩ + by_cases hq : q = Quotient.mk'' g + · exact hq.symm ▸ congr_arg _ (Function.update_self (Quotient.mk'' g) g Quotient.out) + · refine Function.update_of_ne ?_ g Quotient.out ▸ q.out_eq' + exact hq + +set_option linter.deprecated false in +@[to_additive (attr := deprecated exists_isComplement_right (since := "2024-12-18"))] lemma exists_right_transversal (H : Subgroup G) (g : G) : ∃ S ∈ rightTransversals (H : Set G), g ∈ S := by classical @@ -282,7 +391,7 @@ lemma exists_left_transversal_of_le {H' H : Subgroup G} (h : H' ≤ H) : let H'' : Subgroup H := H'.comap H.subtype have : H' = H''.map H.subtype := by simp [H'', h] rw [this] - obtain ⟨S, cmem, -⟩ := H''.exists_left_transversal 1 + obtain ⟨S, cmem, -⟩ := H''.exists_isComplement_left 1 refine ⟨H.subtype '' S, ?_, ?_⟩ · have : H.subtype '' (S * H'') = H.subtype '' S * H''.map H.subtype := image_mul H.subtype rw [← this, cmem.mul_eq] @@ -298,7 +407,7 @@ lemma exists_right_transversal_of_le {H' H : Subgroup G} (h : H' ≤ H) : let H'' : Subgroup H := H'.comap H.subtype have : H' = H''.map H.subtype := by simp [H'', h] rw [this] - obtain ⟨S, cmem, -⟩ := H''.exists_right_transversal 1 + obtain ⟨S, cmem, -⟩ := H''.exists_isComplement_right 1 refine ⟨H.subtype '' S, ?_, ?_⟩ · have : H.subtype '' (H'' * S) = H''.map H.subtype * H.subtype '' S := image_mul H.subtype rw [← this, cmem.mul_eq] @@ -337,7 +446,7 @@ theorem equiv_fst_eq_iff_leftCosetEquivalence {g₁ g₂ : G} : mul_inv_rev, ← mul_assoc, inv_mul_cancel_right, ← coe_inv, ← coe_mul] exact Subtype.property _ · intro h - apply (mem_leftTransversals_iff_existsUnique_inv_mul_mem.1 hSK g₁).unique + apply (isComplement_iff_existsUnique_inv_mul_mem.1 hSK g₁).unique · -- This used to be `simp [...]` before https://github.com/leanprover/lean4/pull/2644 rw [equiv_fst_eq_mul_inv]; simp · rw [SetLike.mem_coe, ← mul_mem_cancel_right h] @@ -353,7 +462,7 @@ theorem equiv_snd_eq_iff_rightCosetEquivalence {g₁ g₂ : G} : mul_inv_rev, mul_assoc, mul_inv_cancel_left, ← coe_inv, ← coe_mul] exact Subtype.property _ · intro h - apply (mem_rightTransversals_iff_existsUnique_mul_inv_mem.1 hHT g₁).unique + apply (isComplement_iff_existsUnique_mul_inv_mem.1 hHT g₁).unique · -- This used to be `simp [...]` before https://github.com/leanprover/lean4/pull/2644 rw [equiv_snd_eq_inv_mul]; simp · rw [SetLike.mem_coe, ← mul_mem_cancel_left h] @@ -451,111 +560,149 @@ theorem coe_equiv_snd_eq_one_iff_mem {g : G} (h1 : 1 ∈ T) : ((hST.equiv g).snd : G) = 1 ↔ g ∈ S := by rw [equiv_snd_eq_inv_mul, inv_mul_eq_one, equiv_fst_eq_self_iff_mem _ h1] -end IsComplement - -namespace MemLeftTransversals - /-- A left transversal is in bijection with left cosets. -/ @[to_additive "A left transversal is in bijection with left cosets."] -noncomputable def toEquiv (hS : S ∈ Subgroup.leftTransversals (H : Set G)) : G ⧸ H ≃ S := - (Equiv.ofBijective _ (Subgroup.mem_leftTransversals_iff_bijective.mp hS)).symm - -@[to_additive "A left transversal is finite iff the subgroup has finite index"] -theorem finite_iff - (h : S ∈ Subgroup.leftTransversals (H : Set G)) : - Finite S ↔ H.FiniteIndex := by - rw [← (Subgroup.MemLeftTransversals.toEquiv h).finite_iff] +noncomputable def leftQuotientEquiv (hS : IsComplement S H) : G ⧸ H ≃ S := + (Equiv.ofBijective _ (isComplement_subgroup_right_iff_bijective.mp hS)).symm + +@[deprecated (since := "2024-12-28")] +alias _root_.Subgroup.MemLeftTransversals.toEquiv := leftQuotientEquiv + +/-- A left transversal is finite iff the subgroup has finite index.-/ +@[to_additive "A left transversal is finite iff the subgroup has finite index."] +theorem finite_left_iff (h : IsComplement S H) : Finite S ↔ H.FiniteIndex := by + rw [← h.leftQuotientEquiv.finite_iff] exact ⟨fun _ ↦ finiteIndex_of_finite_quotient H, fun _ ↦ finite_quotient_of_finiteIndex H⟩ +@[deprecated (since := "2024-12-28")] +alias _root_.Subgroup.MemLeftTransversals.finite_iff := finite_left_iff + @[to_additive] -theorem mk''_toEquiv (hS : S ∈ Subgroup.leftTransversals (H : Set G)) (q : G ⧸ H) : - Quotient.mk'' (toEquiv hS q : G) = q := - (toEquiv hS).symm_apply_apply q +theorem quotientGroupMk_leftQuotientEquiv (hS : IsComplement S H) (q : G ⧸ H) : + Quotient.mk'' (leftQuotientEquiv hS q : G) = q := + hS.leftQuotientEquiv.symm_apply_apply q + +@[deprecated (since := "2024-12-28")] +alias _root_.Subgroup.MemLeftTransversals.mk''_toEquiv := quotientGroupMk_leftQuotientEquiv @[to_additive] -theorem toEquiv_apply {f : G ⧸ H → G} (hf : ∀ q, (f q : G ⧸ H) = q) (q : G ⧸ H) : - (toEquiv (range_mem_leftTransversals hf) q : G) = f q := by +theorem leftQuotientEquiv_apply {f : G ⧸ H → G} (hf : ∀ q, (f q : G ⧸ H) = q) (q : G ⧸ H) : + (leftQuotientEquiv (isComplement_range_left hf) q : G) = f q := by refine (Subtype.ext_iff.mp ?_).trans (Subtype.coe_mk (f q) ⟨q, rfl⟩) - exact (toEquiv (range_mem_leftTransversals hf)).apply_eq_iff_eq_symm_apply.mpr (hf q).symm + exact (leftQuotientEquiv (isComplement_range_left hf)).apply_eq_iff_eq_symm_apply.mpr (hf q).symm + +@[deprecated (since := "2024-12-28")] +alias _root_.Subgroup.MemLeftTransversals.toEquiv_apply := leftQuotientEquiv_apply /-- A left transversal can be viewed as a function mapping each element of the group to the chosen representative from that left coset. -/ @[to_additive "A left transversal can be viewed as a function mapping each element of the group to the chosen representative from that left coset."] -noncomputable def toFun (hS : S ∈ Subgroup.leftTransversals (H : Set G)) : G → S := - toEquiv hS ∘ Quotient.mk'' +noncomputable def toLeftFun (hS : IsComplement S H) : G → S := leftQuotientEquiv hS ∘ Quotient.mk'' -@[to_additive] -theorem inv_toFun_mul_mem (hS : S ∈ Subgroup.leftTransversals (H : Set G)) (g : G) : - (toFun hS g : G)⁻¹ * g ∈ H := - QuotientGroup.leftRel_apply.mp <| Quotient.exact' <| mk''_toEquiv _ _ +@[deprecated (since := "2024-12-28")] +alias _root_.Subgroup.MemLeftTransversals.toFun := toLeftFun @[to_additive] -theorem inv_mul_toFun_mem (hS : S ∈ Subgroup.leftTransversals (H : Set G)) (g : G) : - g⁻¹ * toFun hS g ∈ H := - (congr_arg (· ∈ H) (by rw [mul_inv_rev, inv_inv])).mp (H.inv_mem (inv_toFun_mul_mem hS g)) +theorem inv_toLeftFun_mul_mem (hS : IsComplement S H) (g : G) : + (toLeftFun hS g : G)⁻¹ * g ∈ H := + QuotientGroup.leftRel_apply.mp <| Quotient.exact' <| quotientGroupMk_leftQuotientEquiv _ _ -end MemLeftTransversals +@[deprecated (since := "2024-12-28")] +alias _root_.Subgroup.MemLeftTransversals.inv_toFun_mul_mem := inv_toLeftFun_mul_mem -namespace MemRightTransversals +@[to_additive] +theorem inv_mul_toLeftFun_mem (hS : IsComplement S H) (g : G) : + g⁻¹ * toLeftFun hS g ∈ H := + (congr_arg (· ∈ H) (by rw [mul_inv_rev, inv_inv])).mp (H.inv_mem (inv_toLeftFun_mul_mem hS g)) + +@[deprecated (since := "2024-12-28")] +alias _root_.Subgroup.MemLeftTransversals.inv_mul_toFun_mem := inv_mul_toLeftFun_mem /-- A right transversal is in bijection with right cosets. -/ @[to_additive "A right transversal is in bijection with right cosets."] -noncomputable def toEquiv (hS : S ∈ Subgroup.rightTransversals (H : Set G)) : - Quotient (QuotientGroup.rightRel H) ≃ S := - (Equiv.ofBijective _ (Subgroup.mem_rightTransversals_iff_bijective.mp hS)).symm - -@[to_additive "A right transversal is finite iff the subgroup has finite index"] -theorem finite_iff - (h : S ∈ Subgroup.rightTransversals (H : Set G)) : - Finite S ↔ H.FiniteIndex := by - rw [← (Subgroup.MemRightTransversals.toEquiv h).finite_iff, +noncomputable def rightQuotientEquiv (hT : IsComplement H T) : + Quotient (QuotientGroup.rightRel H) ≃ T := + (Equiv.ofBijective _ (isComplement_subgroup_left_iff_bijective.mp hT)).symm + +@[deprecated (since := "2024-12-28")] +alias _root_.Subgroup.MemRightTransversals.toEquiv := rightQuotientEquiv + +/-- A right transversal is finite iff the subgroup has finite index. -/ +@[to_additive "A right transversal is finite iff the subgroup has finite index."] +theorem finite_right_iff (h : IsComplement H T) : Finite T ↔ H.FiniteIndex := by + rw [← h.rightQuotientEquiv.finite_iff, (QuotientGroup.quotientRightRelEquivQuotientLeftRel H).finite_iff] exact ⟨fun _ ↦ finiteIndex_of_finite_quotient H, fun _ ↦ finite_quotient_of_finiteIndex H⟩ +@[deprecated (since := "2024-12-28")] +alias _root_.Subgroup.MemRightTransversals.finite_iff := finite_right_iff + @[to_additive] -theorem mk''_toEquiv (hS : S ∈ Subgroup.rightTransversals (H : Set G)) - (q : Quotient (QuotientGroup.rightRel H)) : Quotient.mk'' (toEquiv hS q : G) = q := - (toEquiv hS).symm_apply_apply q +theorem mk''_rightQuotientEquiv (hT : IsComplement H T) + (q : Quotient (QuotientGroup.rightRel H)) : Quotient.mk'' (rightQuotientEquiv hT q : G) = q := + (rightQuotientEquiv hT).symm_apply_apply q + +@[deprecated (since := "2024-12-28")] +alias _root_.Subgroup.MemRightTransversals.mk''_toEquiv := mk''_rightQuotientEquiv @[to_additive] -theorem toEquiv_apply {f : Quotient (QuotientGroup.rightRel H) → G} +theorem rightQuotientEquiv_apply {f : Quotient (QuotientGroup.rightRel H) → G} (hf : ∀ q, Quotient.mk'' (f q) = q) (q : Quotient (QuotientGroup.rightRel H)) : - (toEquiv (range_mem_rightTransversals hf) q : G) = f q := by + (rightQuotientEquiv (isComplement_range_right hf) q : G) = f q := by refine (Subtype.ext_iff.mp ?_).trans (Subtype.coe_mk (f q) ⟨q, rfl⟩) - exact (toEquiv (range_mem_rightTransversals hf)).apply_eq_iff_eq_symm_apply.mpr (hf q).symm + exact (rightQuotientEquiv (isComplement_range_right hf)).apply_eq_iff_eq_symm_apply.2 (hf q).symm + +@[deprecated (since := "2024-12-28")] +alias _root_.Subgroup.MemRightTransversals.toEquiv_apply := rightQuotientEquiv_apply /-- A right transversal can be viewed as a function mapping each element of the group to the chosen representative from that right coset. -/ @[to_additive "A right transversal can be viewed as a function mapping each element of the group to the chosen representative from that right coset."] -noncomputable def toFun (hS : S ∈ Subgroup.rightTransversals (H : Set G)) : G → S := - toEquiv hS ∘ Quotient.mk'' +noncomputable def toRightFun (hT : IsComplement H T) : G → T := rightQuotientEquiv hT ∘ .mk'' + +@[deprecated (since := "2024-12-28")] +alias _root_.Subgroup.MemRightTransversals.toFun := toRightFun @[to_additive] -theorem mul_inv_toFun_mem (hS : S ∈ Subgroup.rightTransversals (H : Set G)) (g : G) : - g * (toFun hS g : G)⁻¹ ∈ H := - QuotientGroup.rightRel_apply.mp <| Quotient.exact' <| mk''_toEquiv _ _ +theorem mul_inv_toRightFun_mem (hT : IsComplement H T) (g : G) : + g * (toRightFun hT g : G)⁻¹ ∈ H := + QuotientGroup.rightRel_apply.mp <| Quotient.exact' <| mk''_rightQuotientEquiv _ _ + +@[deprecated (since := "2024-12-28")] +alias _root_.Subgroup.MemRighTransversals.mul_inv_toFun_mem := mul_inv_toRightFun_mem @[to_additive] -theorem toFun_mul_inv_mem (hS : S ∈ Subgroup.rightTransversals (H : Set G)) (g : G) : - (toFun hS g : G) * g⁻¹ ∈ H := - (congr_arg (· ∈ H) (by rw [mul_inv_rev, inv_inv])).mp (H.inv_mem (mul_inv_toFun_mem hS g)) +theorem toRightFun_mul_inv_mem (hT : IsComplement H T) (g : G) : + (toRightFun hT g : G) * g⁻¹ ∈ H := + (congr_arg (· ∈ H) (by rw [mul_inv_rev, inv_inv])).mp (H.inv_mem (mul_inv_toRightFun_mem hT g)) -end MemRightTransversals +@[deprecated (since := "2024-12-28")] +alias _root_.Subgroup.MemRighTransversals.toFun_mul_inv_mem := toRightFun_mul_inv_mem + +end IsComplement section Action open Pointwise MulAction MemLeftTransversals +/-- The collection of left transversals of a subgroup.-/ +@[to_additive "The collection of left transversals of a subgroup."] +abbrev LeftTransversal (H : Subgroup G) := {S : Set G // IsComplement S H} + +/-- The collection of right transversals of a subgroup.-/ +@[to_additive "The collection of right transversals of a subgroup."] +abbrev RightTransversal (H : Subgroup G) := {T : Set G // IsComplement H T} + variable {F : Type*} [Group F] [MulAction F G] [QuotientAction F H] @[to_additive] -noncomputable instance : MulAction F (leftTransversals (H : Set G)) where +noncomputable instance : MulAction F H.LeftTransversal where smul f T := ⟨f • (T : Set G), by - refine mem_leftTransversals_iff_existsUnique_inv_mul_mem.mpr fun g => ?_ - obtain ⟨t, ht1, ht2⟩ := mem_leftTransversals_iff_existsUnique_inv_mul_mem.mp T.2 (f⁻¹ • g) + refine isComplement_iff_existsUnique_inv_mul_mem.mpr fun g => ?_ + obtain ⟨t, ht1, ht2⟩ := isComplement_iff_existsUnique_inv_mul_mem.mp T.2 (f⁻¹ • g) refine ⟨⟨f • (t : G), Set.smul_mem_smul_set t.2⟩, ?_, ?_⟩ · exact smul_inv_smul f g ▸ QuotientAction.inv_mul_mem f ht1 · rintro ⟨-, t', ht', rfl⟩ h @@ -566,32 +713,34 @@ noncomputable instance : MulAction F (leftTransversals (H : Set G)) where mul_smul f₁ f₂ T := Subtype.ext (mul_smul f₁ f₂ (T : Set G)) @[to_additive] -theorem smul_toFun (f : F) (T : leftTransversals (H : Set G)) (g : G) : - (f • (toFun T.2 g : G)) = toFun (f • T).2 (f • g) := - Subtype.ext_iff.mp <| @ExistsUnique.unique (↥(f • (T : Set G))) (fun s => (↑s)⁻¹ * f • g ∈ H) - (mem_leftTransversals_iff_existsUnique_inv_mul_mem.mp (f • T).2 (f • g)) - ⟨f • (toFun T.2 g : G), Set.smul_mem_smul_set (Subtype.coe_prop _)⟩ (toFun (f • T).2 (f • g)) - (QuotientAction.inv_mul_mem f (inv_toFun_mul_mem T.2 g)) (inv_toFun_mul_mem (f • T).2 (f • g)) +theorem smul_toLeftFun (f : F) (S : H.LeftTransversal) (g : G) : + (f • (S.2.toLeftFun g : G)) = (f • S).2.toLeftFun (f • g) := + Subtype.ext_iff.mp <| @ExistsUnique.unique (↥(f • (S : Set G))) (fun s => (↑s)⁻¹ * f • g ∈ H) + (isComplement_iff_existsUnique_inv_mul_mem.mp (f • S).2 (f • g)) + ⟨f • (S.2.toLeftFun g : G), Set.smul_mem_smul_set (Subtype.coe_prop _)⟩ + ((f • S).2.toLeftFun (f • g)) + (QuotientAction.inv_mul_mem f (S.2.inv_toLeftFun_mul_mem g)) + ((f • S).2.inv_toLeftFun_mul_mem (f • g)) @[to_additive] -theorem smul_toEquiv (f : F) (T : leftTransversals (H : Set G)) (q : G ⧸ H) : - f • (toEquiv T.2 q : G) = toEquiv (f • T).2 (f • q) := - Quotient.inductionOn' q fun g => smul_toFun f T g +theorem smul_leftQuotientEquiv (f : F) (S : H.LeftTransversal) (q : G ⧸ H) : + f • (S.2.leftQuotientEquiv q : G) = (f • S).2.leftQuotientEquiv (f • q) := + Quotient.inductionOn' q fun g => smul_toLeftFun f S g @[to_additive] -theorem smul_apply_eq_smul_apply_inv_smul (f : F) (T : leftTransversals (H : Set G)) (q : G ⧸ H) : - (toEquiv (f • T).2 q : G) = f • (toEquiv T.2 (f⁻¹ • q) : G) := by - rw [smul_toEquiv, smul_inv_smul] +theorem smul_apply_eq_smul_apply_inv_smul (f : F) (S : H.LeftTransversal) (q : G ⧸ H) : + ((f • S).2.leftQuotientEquiv q : G) = f • (S.2.leftQuotientEquiv (f⁻¹ • q) : G) := by + rw [smul_leftQuotientEquiv, smul_inv_smul] end Action @[to_additive] -instance : Inhabited (leftTransversals (H : Set G)) := - ⟨⟨Set.range Quotient.out, range_mem_leftTransversals Quotient.out_eq'⟩⟩ +instance : Inhabited H.LeftTransversal := + ⟨⟨Set.range Quotient.out, isComplement_range_left Quotient.out_eq'⟩⟩ @[to_additive] -instance : Inhabited (rightTransversals (H : Set G)) := - ⟨⟨Set.range Quotient.out, range_mem_rightTransversals Quotient.out_eq'⟩⟩ +instance : Inhabited H.RightTransversal := + ⟨⟨Set.range Quotient.out, isComplement_range_right Quotient.out_eq'⟩⟩ theorem IsComplement'.isCompl (h : IsComplement' H K) : IsCompl H K := by refine @@ -611,14 +760,14 @@ theorem IsComplement'.disjoint (h : IsComplement' H K) : Disjoint H K := h.isCompl.disjoint theorem IsComplement'.index_eq_card (h : IsComplement' H K) : K.index = Nat.card H := - (card_left_transversal h).symm + h.card_left.symm /-- If `H` and `K` are complementary with `K` normal, then `G ⧸ K` is isomorphic to `H`. -/ @[simps!] noncomputable def IsComplement'.QuotientMulEquiv [K.Normal] (h : H.IsComplement' K) : G ⧸ K ≃* H := MulEquiv.symm - { (MemLeftTransversals.toEquiv h).symm with + { h.leftQuotientEquiv.symm with map_mul' := fun _ _ ↦ rfl } theorem IsComplement.card_mul (h : IsComplement S T) : diff --git a/Mathlib/GroupTheory/CosetCover.lean b/Mathlib/GroupTheory/CosetCover.lean index 589e4d095bb97..434d749f050b9 100644 --- a/Mathlib/GroupTheory/CosetCover.lean +++ b/Mathlib/GroupTheory/CosetCover.lean @@ -54,10 +54,10 @@ section leftCoset_cover_const theorem exists_leftTransversal_of_FiniteIndex {D H : Subgroup G} [D.FiniteIndex] (hD_le_H : D ≤ H) : ∃ t : Finset H, - (t : Set H) ∈ leftTransversals (D.subgroupOf H) ∧ + IsComplement (t : Set H) (D.subgroupOf H) ∧ ⋃ g ∈ t, (g : G) • (D : Set G) = H := by - have ⟨t, ht⟩ := exists_left_transversal (D.subgroupOf H) 1 - have hf : t.Finite := (MemLeftTransversals.finite_iff ht.1).mpr inferInstance + have ⟨t, ht⟩ := (D.subgroupOf H).exists_isComplement_left 1 + have hf : t.Finite := ht.1.finite_left_iff.mpr inferInstance refine ⟨hf.toFinset, hf.coe_toFinset.symm ▸ ht.1, ?_⟩ ext x suffices (∃ y ∈ t, ∃ d ∈ D, y * d = x) ↔ x ∈ H by simpa using this @@ -65,8 +65,8 @@ theorem exists_leftTransversal_of_FiniteIndex · rintro ⟨⟨y, hy⟩, -, d, h, rfl⟩ exact H.mul_mem hy (hD_le_H h) · intro hx - exact ⟨_, (MemLeftTransversals.toFun ht.1 ⟨x, hx⟩).2, _, - MemLeftTransversals.inv_toFun_mul_mem ht.1 ⟨x, hx⟩, mul_inv_cancel_left _ _⟩ + exact ⟨_, (ht.1.toLeftFun ⟨x, hx⟩).2, _, + ht.1.inv_toLeftFun_mul_mem ⟨x, hx⟩, mul_inv_cancel_left _ _⟩ variable {ι : Type*} {s : Finset ι} {H : Subgroup G} {g : ι → G} @@ -238,7 +238,7 @@ theorem leftCoset_cover_filter_FiniteIndex_aux have ⟨k, hkfi, hk⟩ : ∃ k, (H k.1.1).FiniteIndex ∧ K k = D := have ⟨j, hj, hjfi⟩ := exists_finiteIndex_of_leftCoset_cover hcovers have ⟨x, hx⟩ : (t j hj hjfi).Nonempty := Finset.nonempty_coe_sort.mp - (MemLeftTransversals.toEquiv (ht j hj hjfi).1).symm.nonempty + (ht j hj hjfi).1.leftQuotientEquiv.symm.nonempty ⟨⟨⟨j, hj⟩, ⟨x, dif_pos hjfi ▸ hx⟩⟩, hjfi, if_pos hjfi⟩ -- Since `D` is the unique subgroup of finite index whose cosets occur in the new covering, -- the cosets of the other subgroups can be omitted. @@ -260,7 +260,7 @@ theorem leftCoset_cover_filter_FiniteIndex_aux by_cases hfi : (H i).FiniteIndex · rw [← relindex_mul_index (hD_le i.2 hfi), Nat.cast_mul, mul_comm, mul_inv_cancel_right₀ (Nat.cast_ne_zero.mpr hfi.finiteIndex)] - simpa [K, hfi] using card_left_transversal (ht i.1 i.2 hfi).1 + simpa [K, hfi] using (ht i.1 i.2 hfi).1.card_left · rw [of_not_not (FiniteIndex.mk.mt hfi), Nat.cast_zero, inv_zero, zero_mul] simpa [K, hfi] using hHD i hfi refine ⟨?_, ?_, ?_⟩ diff --git a/Mathlib/GroupTheory/HNNExtension.lean b/Mathlib/GroupTheory/HNNExtension.lean index 0134476792c78..def2ac3cf71c4 100644 --- a/Mathlib/GroupTheory/HNNExtension.lean +++ b/Mathlib/GroupTheory/HNNExtension.lean @@ -181,7 +181,7 @@ structure TransversalPair : Type _ where compl : ∀ u, IsComplement (toSubgroup A B u : Subgroup G) (set u) instance TransversalPair.nonempty : Nonempty (TransversalPair G A B) := by - choose t ht using fun u ↦ (toSubgroup A B u).exists_right_transversal 1 + choose t ht using fun u ↦ (toSubgroup A B u).exists_isComplement_right 1 exact ⟨⟨t, fun i ↦ (ht i).1⟩⟩ /-- A reduced word is a `head`, which is an element of `G`, followed by the product list of pairs. diff --git a/Mathlib/GroupTheory/PushoutI.lean b/Mathlib/GroupTheory/PushoutI.lean index 9c027378a289c..e774edda3b70a 100644 --- a/Mathlib/GroupTheory/PushoutI.lean +++ b/Mathlib/GroupTheory/PushoutI.lean @@ -228,7 +228,7 @@ structure Transversal : Type _ where compl : ∀ i, IsComplement (φ i).range (set i) theorem transversal_nonempty (hφ : ∀ i, Injective (φ i)) : Nonempty (Transversal φ) := by - choose t ht using fun i => (φ i).range.exists_right_transversal 1 + choose t ht using fun i => (φ i).range.exists_isComplement_right 1 apply Nonempty.intro exact { injective := hφ diff --git a/Mathlib/GroupTheory/Schreier.lean b/Mathlib/GroupTheory/Schreier.lean index a93ab6728cfdc..d7bbfc3d60879 100644 --- a/Mathlib/GroupTheory/Schreier.lean +++ b/Mathlib/GroupTheory/Schreier.lean @@ -15,7 +15,7 @@ In this file we prove Schreier's lemma. - `closure_mul_image_eq` : **Schreier's Lemma**: If `R : Set G` is a right_transversal of `H : Subgroup G` with `1 ∈ R`, and if `G` is generated by `S : Set G`, - then `H` is generated by the `Set` `(R * S).image (fun g ↦ g * (toFun hR g)⁻¹)`. + then `H` is generated by the `Set` `(R * S).image (fun g ↦ g * (hR.toRightFun g)⁻¹)`. - `fg_of_index_ne_zero` : **Schreier's Lemma**: A finite index subgroup of a finitely generated group is finitely generated. - `card_commutator_le_of_finite_commutatorSet`: A theorem of Schur: The size of the commutator @@ -62,9 +62,9 @@ open MemRightTransversals variable {G : Type*} [Group G] {H : Subgroup G} {R S : Set G} theorem closure_mul_image_mul_eq_top - (hR : R ∈ rightTransversals (H : Set G)) (hR1 : (1 : G) ∈ R) (hS : closure S = ⊤) : - (closure ((R * S).image fun g => g * (toFun hR g : G)⁻¹)) * R = ⊤ := by - let f : G → R := fun g => toFun hR g + (hR : IsComplement H R) (hR1 : (1 : G) ∈ R) (hS : closure S = ⊤) : + (closure ((R * S).image fun g => g * (hR.toRightFun g : G)⁻¹)) * R = ⊤ := by + let f : G → R := hR.toRightFun let U : Set G := (R * S).image fun g => g * (f g : G)⁻¹ change (closure U : Set G) * R = ⊤ refine top_le_iff.mp fun g _ => ?_ @@ -79,26 +79,26 @@ theorem closure_mul_image_mul_eq_top refine Set.mul_mem_mul ((closure U).mul_mem hu ((closure U).inv_mem ?_)) (f (r * s⁻¹)).2 refine subset_closure ⟨f (r * s⁻¹) * s, Set.mul_mem_mul (f (r * s⁻¹)).2 hs, ?_⟩ rw [mul_right_inj, inv_inj, ← Subtype.coe_mk r hr, ← Subtype.ext_iff, Subtype.coe_mk] - apply (mem_rightTransversals_iff_existsUnique_mul_inv_mem.mp hR (f (r * s⁻¹) * s)).unique - (mul_inv_toFun_mem hR (f (r * s⁻¹) * s)) + apply (isComplement_iff_existsUnique_mul_inv_mem.mp hR (f (r * s⁻¹) * s)).unique + (hR.mul_inv_toRightFun_mem (f (r * s⁻¹) * s)) rw [mul_assoc, ← inv_inv s, ← mul_inv_rev, inv_inv] - exact toFun_mul_inv_mem hR (r * s⁻¹) + exact hR.toRightFun_mul_inv_mem (r * s⁻¹) /-- **Schreier's Lemma**: If `R : Set G` is a `rightTransversal` of `H : Subgroup G` with `1 ∈ R`, and if `G` is generated by `S : Set G`, then `H` is generated by the `Set` - `(R * S).image (fun g ↦ g * (toFun hR g)⁻¹)`. -/ -theorem closure_mul_image_eq (hR : R ∈ rightTransversals (H : Set G)) (hR1 : (1 : G) ∈ R) - (hS : closure S = ⊤) : closure ((R * S).image fun g => g * (toFun hR g : G)⁻¹) = H := by - have hU : closure ((R * S).image fun g => g * (toFun hR g : G)⁻¹) ≤ H := by + `(R * S).image (fun g ↦ g * (hR.toRightFun g)⁻¹)`. -/ +theorem closure_mul_image_eq (hR : IsComplement H R) (hR1 : (1 : G) ∈ R) + (hS : closure S = ⊤) : closure ((R * S).image fun g => g * (hR.toRightFun g : G)⁻¹) = H := by + have hU : closure ((R * S).image fun g => g * (hR.toRightFun g : G)⁻¹) ≤ H := by rw [closure_le] rintro - ⟨g, -, rfl⟩ - exact mul_inv_toFun_mem hR g + exact hR.mul_inv_toRightFun_mem g refine le_antisymm hU fun h hh => ?_ obtain ⟨g, hg, r, hr, rfl⟩ := show h ∈ _ from eq_top_iff.mp (closure_mul_image_mul_eq_top hR hR1 hS) (mem_top h) 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 + apply (isComplement_iff_existsUnique_mul_inv_mem.mp hR r).unique · rw [Subtype.coe_mk, mul_inv_cancel] exact H.one_mem · rw [Subtype.coe_mk, inv_one, mul_one] @@ -106,20 +106,20 @@ theorem closure_mul_image_eq (hR : R ∈ rightTransversals (H : Set G)) (hR1 : ( /-- **Schreier's Lemma**: If `R : Set G` is a `rightTransversal` of `H : Subgroup G` with `1 ∈ R`, and if `G` is generated by `S : Set G`, then `H` is generated by the `Set` - `(R * S).image (fun g ↦ g * (toFun hR g)⁻¹)`. -/ -theorem closure_mul_image_eq_top (hR : R ∈ rightTransversals (H : Set G)) (hR1 : (1 : G) ∈ R) + `(R * S).image (fun g ↦ g * (hR.toRightFun g)⁻¹)`. -/ +theorem closure_mul_image_eq_top (hR : IsComplement H R) (hR1 : (1 : G) ∈ R) (hS : closure S = ⊤) : closure ((R * S).image fun g => - ⟨g * (toFun hR g : G)⁻¹, mul_inv_toFun_mem hR g⟩ : Set H) = ⊤ := by + ⟨g * (hR.toRightFun g : G)⁻¹, hR.mul_inv_toRightFun_mem g⟩ : Set H) = ⊤ := by rw [eq_top_iff, ← map_subtype_le_map_subtype, MonoidHom.map_closure, Set.image_image] exact (map_subtype_le ⊤).trans (ge_of_eq (closure_mul_image_eq hR hR1 hS)) /-- **Schreier's Lemma**: If `R : Finset G` is a `rightTransversal` of `H : Subgroup G` with `1 ∈ R`, and if `G` is generated by `S : Finset G`, then `H` is generated by the `Finset` - `(R * S).image (fun g ↦ g * (toFun hR g)⁻¹)`. -/ + `(R * S).image (fun g ↦ g * (hR.toRightFun g)⁻¹)`. -/ theorem closure_mul_image_eq_top' [DecidableEq G] {R S : Finset G} - (hR : (R : Set G) ∈ rightTransversals (H : Set G)) (hR1 : (1 : G) ∈ R) + (hR : IsComplement (H : Set G) R) (hR1 : (1 : G) ∈ R) (hS : closure (S : Set G) = ⊤) : - closure (((R * S).image fun g => ⟨_, mul_inv_toFun_mem hR g⟩ : Finset H) : Set H) = ⊤ := by + closure (((R * S).image fun g => ⟨_, hR.mul_inv_toRightFun_mem g⟩ : Finset H) : Set H) = ⊤ := by rw [Finset.coe_image, Finset.coe_mul] exact closure_mul_image_eq_top hR hR1 hS @@ -129,10 +129,10 @@ theorem exists_finset_card_le_mul [FiniteIndex H] {S : Finset G} (hS : closure ( ∃ T : Finset H, T.card ≤ H.index * S.card ∧ closure (T : Set H) = ⊤ := by letI := H.fintypeQuotientOfFiniteIndex haveI : DecidableEq G := Classical.decEq G - obtain ⟨R₀, hR, hR1⟩ := H.exists_right_transversal 1 - haveI : Fintype R₀ := Fintype.ofEquiv _ (toEquiv hR) + obtain ⟨R₀, hR, hR1⟩ := H.exists_isComplement_right 1 + haveI : Fintype R₀ := Fintype.ofEquiv _ hR.rightQuotientEquiv let R : Finset G := Set.toFinset R₀ - replace hR : (R : Set G) ∈ rightTransversals (H : Set G) := by rwa [Set.coe_toFinset] + replace hR : IsComplement (H : Set G) R := by rwa [Set.coe_toFinset] replace hR1 : (1 : G) ∈ R := by rwa [Set.mem_toFinset] refine ⟨_, ?_, closure_mul_image_eq_top' hR hR1 hS⟩ calc @@ -142,7 +142,7 @@ theorem exists_finset_card_le_mul [FiniteIndex H] {S : Finset G} (hS : closure ( _ = H.index * S.card := congr_arg (· * S.card) ?_ calc R.card = Fintype.card R := (Fintype.card_coe R).symm - _ = _ := (Fintype.card_congr (toEquiv hR)).symm + _ = _ := (Fintype.card_congr hR.rightQuotientEquiv).symm _ = Fintype.card (G ⧸ H) := QuotientGroup.card_quotient_rightRel H _ = H.index := by rw [index_eq_card, Nat.card_eq_fintype_card] diff --git a/Mathlib/GroupTheory/SchurZassenhaus.lean b/Mathlib/GroupTheory/SchurZassenhaus.lean index 2e959aacc20e4..b36b24b2c9fa8 100644 --- a/Mathlib/GroupTheory/SchurZassenhaus.lean +++ b/Mathlib/GroupTheory/SchurZassenhaus.lean @@ -28,7 +28,7 @@ section SchurZassenhausAbelian open MulOpposite MulAction Subgroup.leftTransversals MemLeftTransversals variable {G : Type*} [Group G] (H : Subgroup G) [IsCommutative H] [FiniteIndex H] - (α β : leftTransversals (H : Set G)) + (α β : H.LeftTransversal) /-- The quotient of the transversals of an abelian normal `N` by the `diff` relation. -/ def QuotientDiff := diff --git a/Mathlib/GroupTheory/Transfer.lean b/Mathlib/GroupTheory/Transfer.lean index bd55dae349bfd..f33f7fc8e19a5 100644 --- a/Mathlib/GroupTheory/Transfer.lean +++ b/Mathlib/GroupTheory/Transfer.lean @@ -35,15 +35,15 @@ open Finset MulAction open scoped Pointwise -variable (R S T : leftTransversals (H : Set G)) [FiniteIndex H] +variable (R S T : H.LeftTransversal) [FiniteIndex H] /-- The difference of two left transversals -/ @[to_additive "The difference of two left transversals"] noncomputable def diff : A := - let α := MemLeftTransversals.toEquiv S.2 - let β := MemLeftTransversals.toEquiv T.2 - (@Finset.univ (G ⧸ H) H.fintypeQuotientOfFiniteIndex).prod fun q => - ϕ + let α := S.2.leftQuotientEquiv + let β := T.2.leftQuotientEquiv + let _ := H.fintypeQuotientOfFiniteIndex + ∏ q : G ⧸ H, ϕ ⟨(α q : G)⁻¹ * β q, QuotientGroup.leftRel_apply.mp <| Quotient.exact' ((α.symm_apply_apply q).trans (β.symm_apply_apply q).symm)⟩ @@ -103,23 +103,23 @@ lemma mem_transferSet (q : G ⧸ H) : transferFunction H g q ∈ transferSet H g variable (H) in /-- The transfer transversal. Contains elements of the form `g ^ k • g₀` for fixed choices of representatives `g₀` of fixed choices of representatives `q₀` of `⟨g⟩`-orbits in `G ⧸ H`. -/ -def transferTransversal : leftTransversals (H : Set G) := - ⟨transferSet H g, range_mem_leftTransversals (coe_transferFunction g)⟩ +def transferTransversal : H.LeftTransversal := + ⟨transferSet H g, isComplement_range_left (coe_transferFunction g)⟩ lemma transferTransversal_apply (q : G ⧸ H) : - ↑(toEquiv (transferTransversal H g).2 q) = transferFunction H g q := - toEquiv_apply (coe_transferFunction g) q + ↑((transferTransversal H g).2.leftQuotientEquiv q) = transferFunction H g q := + IsComplement.leftQuotientEquiv_apply (coe_transferFunction g) q lemma transferTransversal_apply' (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) (k : ZMod (minimalPeriod (g • ·) q.out)) : - ↑(toEquiv (transferTransversal H g).2 (g ^ (cast k : ℤ) • q.out)) = + ↑((transferTransversal H g).2.leftQuotientEquiv (g ^ (cast k : ℤ) • q.out)) = g ^ (cast k : ℤ) * q.out.out := by rw [transferTransversal_apply, transferFunction_apply, ← quotientEquivSigmaZMod_symm_apply, apply_symm_apply] lemma transferTransversal_apply'' (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) (k : ZMod (minimalPeriod (g • ·) q.out)) : - ↑(toEquiv (g • transferTransversal H g).2 (g ^ (cast k : ℤ) • q.out)) = + ↑((g • transferTransversal H g).2.leftQuotientEquiv (g ^ (cast k : ℤ) • q.out)) = if k = 0 then g ^ minimalPeriod (g • ·) q.out * q.out.out else g ^ (cast k : ℤ) * q.out.out := by rw [smul_apply_eq_smul_apply_inv_smul, transferTransversal_apply, transferFunction_apply, ← @@ -141,14 +141,14 @@ the transfer homomorphism is `transfer ϕ : G →* A`. -/ @[to_additive "Given `ϕ : H →+ A` from `H : AddSubgroup G` to an additive commutative group `A`, the transfer homomorphism is `transfer ϕ : G →+ A`."] noncomputable def transfer [FiniteIndex H] : G →* A := - let T : leftTransversals (H : Set G) := Inhabited.default + let T : H.LeftTransversal := default { toFun := fun g => diff ϕ T (g • T) -- Porting note (https://github.com/leanprover-community/mathlib4/issues/12129): additional beta reduction needed map_one' := by beta_reduce; rw [one_smul, diff_self] -- Porting note: added `simp only` (not just beta reduction) map_mul' := fun g h => by simp only; rw [mul_smul, ← diff_mul_diff, smul_diff_smul] } -variable (T : leftTransversals (H : Set G)) +variable (T : H.LeftTransversal) @[to_additive] theorem transfer_def [FiniteIndex H] (g : G) : transfer ϕ g = diff ϕ T (g • T) := by From e2c77f76b7cac7086908e450cf291c755324d045 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 3 Jan 2025 11:02:36 +0000 Subject: [PATCH 108/189] fix: avoid using custom parser combinators (#20355) These are not supported unless `enableInitializersExecution` is called, and various downstream scripts do not call this. https://leanprover.zulipchat.com/#narrow/channel/270676-lean4/topic/lean4checker.20failure/near/486191171 has some context for the underlying issue, which seems to be that Lean tries to finalize environment extensions even if initializers are disabled, leading to crashes if those extensions (such as the parser alias table) expect initializers to have run. To prove this works, this removes `Lean.enableInitializersExecution` from the script to check the yaml files. This is what we will want to do anyway after https://github.com/leanprover/lean4/pull/6325 lands, as this script should not need to run any extensions. --- Mathlib/Analysis/InnerProductSpace/PiL2.lean | 3 ++- Mathlib/Util/Superscript.lean | 25 +++++++++++++++++--- scripts/check-yaml.lean | 1 - 3 files changed, 24 insertions(+), 5 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index e45fd08b5881d..b3d63c949ef5e 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -106,12 +106,13 @@ abbrev EuclideanSpace (𝕜 : Type*) (n : Type*) : Type _ := section Notation open Lean Meta Elab Term Macro TSyntax PrettyPrinter.Delaborator SubExpr +open Mathlib.Tactic (subscriptTerm) /-- Notation for vectors in Lp space. `!₂[x, y, ...]` is a shorthand for `(WithLp.equiv 2 _ _).symm ![x, y, ...]`, of type `EuclideanSpace _ (Fin _)`. This also works for other subscripts. -/ -syntax (name := PiLp.vecNotation) "!" noWs subscript(term) noWs "[" term,* "]" : term +syntax (name := PiLp.vecNotation) "!" noWs subscriptTerm noWs "[" term,* "]" : term macro_rules | `(!$p:subscript[$e:term,*]) => do -- override the `Fin n.succ` to a literal let n := e.getElems.size diff --git a/Mathlib/Util/Superscript.lean b/Mathlib/Util/Superscript.lean index e322df63fe38f..88c027c08dadb 100644 --- a/Mathlib/Util/Superscript.lean +++ b/Mathlib/Util/Superscript.lean @@ -251,6 +251,17 @@ def superscript.parenthesizer := Superscript.scriptParser.parenthesizer ``supers def superscript.formatter := Superscript.scriptParser.formatter "superscript" .superscript ``superscript +/-- Shorthand for `superscript(term)`. + +This is needed because the initializer below does not always run, and if it has not run then +downstream parsers using the combinators will crash. + +See https://leanprover.zulipchat.com/#narrow/channel/270676-lean4/topic/Non-builtin.20parser.20aliases/near/365125476 +for some context. -/ +@[term_parser] +def superscriptTerm := leading_parser (withAnonymousAntiquot := false) superscript termParser + +initialize register_parser_alias superscript /-- The parser `subscript(term)` parses a subscript. Basic usage is: @@ -276,8 +287,16 @@ def subscript.parenthesizer := Superscript.scriptParser.parenthesizer ``subscrip @[combinator_formatter subscript] def subscript.formatter := Superscript.scriptParser.formatter "subscript" .subscript ``subscript -initialize - register_parser_alias superscript - register_parser_alias subscript +/-- Shorthand for `subscript(term)`. + +This is needed because the initializer below does not always run, and if it has not run then +downstream parsers using the combinators will crash. + +See https://leanprover.zulipchat.com/#narrow/channel/270676-lean4/topic/Non-builtin.20parser.20aliases/near/365125476 +for some context. -/ +@[term_parser] +def subscriptTerm := leading_parser (withAnonymousAntiquot := false) subscript termParser + +initialize register_parser_alias subscript end Mathlib.Tactic diff --git a/scripts/check-yaml.lean b/scripts/check-yaml.lean index 37da03fd90436..94e463810868d 100644 --- a/scripts/check-yaml.lean +++ b/scripts/check-yaml.lean @@ -40,7 +40,6 @@ def processDb (decls : ConstMap) : String → IO Bool return false unsafe def main : IO Unit := do - Lean.enableInitializersExecution CoreM.withImportModules #[`Mathlib, `Archive] (searchPath := compile_time_search_path%) (trustLevel := 1024) do let decls := (← getEnv).constants From 150f720a3d74d448a7ffea66217670c1862c9602 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 3 Jan 2025 11:02:37 +0000 Subject: [PATCH 109/189] feat(Algebra/Lie): add `LieTower` typeclass abstracting the Leibniz rule (#20403) This PR is a preparation for follow-up work on Lie's theorem. In the proof of Lie's theorem, there are a number of different ways to infer bracket actions, and not all of them allowed rewriting by the `leibniz_lie` lemma. So the proof would be massaging terms of Lie ideals into terms of the ambient Lie algebra, etc... By abstracting this lemma, and providing suitable instances, we can now simply rewrite with `leibniz_lie` without the massaging. --- Mathlib/Algebra/Lie/Basic.lean | 32 +++++++++++++++++++++++++++-- Mathlib/Algebra/Lie/Subalgebra.lean | 17 +++++++++------ Mathlib/Algebra/Lie/Submodule.lean | 6 ++++++ 3 files changed, 47 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/Lie/Basic.lean b/Mathlib/Algebra/Lie/Basic.lean index 8ff00bfb403c6..8c75a9021fe88 100644 --- a/Mathlib/Algebra/Lie/Basic.lean +++ b/Mathlib/Algebra/Lie/Basic.lean @@ -92,6 +92,34 @@ class LieModule (R : Type u) (L : Type v) (M : Type w) [CommRing R] [LieRing L] /-- A Lie module bracket is compatible with scalar multiplication in its second argument. -/ protected lie_smul : ∀ (t : R) (x : L) (m : M), ⁅x, t • m⁆ = t • ⁅x, m⁆ +/-- A tower of Lie bracket actions encapsulates the Leibniz rule for Lie bracket actions. + +More precisely, it does so in a relative setting: +Let `L₁` and `L₂` be two types with Lie bracket actions on a type `M` endowed with an addition, +and additionally assume a Lie bracket action of `L₁` on `L₂`. +Then the Leibniz rule asserts for all `x : L₁`, `y : L₂`, and `m : M` that +`⁅x, ⁅y, m⁆⁆ = ⁅⁅x, y⁆, m⁆ + ⁅y, ⁅x, m⁆⁆` holds. + +Common examples include the case where `L₁` is a Lie subalgebra of `L₂` +and the case where `L₂` is a Lie ideal of `L₁`. -/ +class IsLieTower (L₁ L₂ M : Type*) [Bracket L₁ L₂] [Bracket L₁ M] [Bracket L₂ M] [Add M] where + protected leibniz_lie (x : L₁) (y : L₂) (m : M) : ⁅x, ⁅y, m⁆⁆ = ⁅⁅x, y⁆, m⁆ + ⁅y, ⁅x, m⁆⁆ + +section IsLieTower + +variable {L₁ L₂ M : Type*} [Bracket L₁ L₂] [Bracket L₁ M] [Bracket L₂ M] + +lemma leibniz_lie [Add M] [IsLieTower L₁ L₂ M] (x : L₁) (y : L₂) (m : M) : + ⁅x, ⁅y, m⁆⁆ = ⁅⁅x, y⁆, m⁆ + ⁅y, ⁅x, m⁆⁆ := IsLieTower.leibniz_lie x y m + +lemma lie_swap_lie [Bracket L₂ L₁] [AddCommGroup M] [IsLieTower L₁ L₂ M] [IsLieTower L₂ L₁ M] + (x : L₁) (y : L₂) (m : M) : ⁅⁅x, y⁆, m⁆ = -⁅⁅y, x⁆, m⁆ := by + have h1 := leibniz_lie x y m + have h2 := leibniz_lie y x m + convert congr($h1.symm - $h2) using 1 <;> simp only [add_sub_cancel_right, sub_add_cancel_right] + +end IsLieTower + section BasicProperties variable {R : Type u} {L : Type v} {M : Type w} {N : Type w₁} @@ -116,8 +144,8 @@ theorem smul_lie : ⁅t • x, m⁆ = t • ⁅x, m⁆ := theorem lie_smul : ⁅x, t • m⁆ = t • ⁅x, m⁆ := LieModule.lie_smul t x m -theorem leibniz_lie : ⁅x, ⁅y, m⁆⁆ = ⁅⁅x, y⁆, m⁆ + ⁅y, ⁅x, m⁆⁆ := - LieRingModule.leibniz_lie x y m +instance : IsLieTower L L M where + leibniz_lie x y m := LieRingModule.leibniz_lie x y m @[simp] theorem lie_zero : ⁅x, 0⁆ = (0 : M) := diff --git a/Mathlib/Algebra/Lie/Subalgebra.lean b/Mathlib/Algebra/Lie/Subalgebra.lean index 5cf5fab92ae25..42f52c2d75d0b 100644 --- a/Mathlib/Algebra/Lie/Subalgebra.lean +++ b/Mathlib/Algebra/Lie/Subalgebra.lean @@ -218,18 +218,23 @@ section LieModule variable {M : Type w} [AddCommGroup M] [LieRingModule L M] 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. -/ -instance lieRingModule : LieRingModule L' M where +instance : Bracket L' M where bracket x m := ⁅(x : L), m⁆ - add_lie x y m := add_lie (x : L) y m - lie_add x y m := lie_add (x : L) y m - leibniz_lie x y m := leibniz_lie (x : L) y m @[simp] theorem coe_bracket_of_module (x : L') (m : M) : ⁅x, m⁆ = ⁅(x : L), m⁆ := rfl +instance : IsLieTower L' L M where + leibniz_lie x y m := leibniz_lie x.val y m + +/-- 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. -/ +instance lieRingModule : LieRingModule L' M where + add_lie x y m := add_lie (x : L) y m + lie_add x y m := lie_add (x : L) y m + leibniz_lie x y m := leibniz_lie x (y : L) m + variable [Module R M] /-- Given a Lie algebra `L` containing a Lie subalgebra `L' ⊆ L`, together with a Lie module `M` of diff --git a/Mathlib/Algebra/Lie/Submodule.lean b/Mathlib/Algebra/Lie/Submodule.lean index f72a8112bb070..7b1a08a9b5d90 100644 --- a/Mathlib/Algebra/Lie/Submodule.lean +++ b/Mathlib/Algebra/Lie/Submodule.lean @@ -273,6 +273,12 @@ theorem LieIdeal.coe_bracket_of_module {R L : Type*} [CommRing R] [LieRing L] [L instance LieIdeal.lieModule (I : LieIdeal R L) : LieModule R I M := LieSubalgebra.lieModule (I : LieSubalgebra R L) +instance (I : LieIdeal R L) : IsLieTower I L M where + leibniz_lie x y m := leibniz_lie x.val y m + +instance (I : LieIdeal R L) : IsLieTower L I M where + leibniz_lie x y m := leibniz_lie x y.val m + end LieIdeal variable {R M} From 2c04291a30f83d51f13f3c469960b7507f7e10b2 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Fri, 3 Jan 2025 11:02:38 +0000 Subject: [PATCH 110/189] feat: A monotone function on a segment belongs to all `L^p` spaces (#20419) Currently, we have the fact that a monotone function on a segment is integrable. This PR generalizes it to all `L^p` spaces (with the same proof), and deduces the integrability result from the general one. From the Carleson project --- .../Function/LocallyIntegrable.lean | 64 +++++++++++++------ .../Function/LpSeminorm/Basic.lean | 2 + .../Function/LpSeminorm/CompareExp.lean | 35 ++++++++++ .../Function/StronglyMeasurable/Basic.lean | 15 +++++ 4 files changed, 97 insertions(+), 19 deletions(-) diff --git a/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean b/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean index 4fbb38fd6c028..d26cd9cc0d9c0 100644 --- a/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean +++ b/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean @@ -485,40 +485,66 @@ open scoped ENNReal section Monotone variable [BorelSpace X] [ConditionallyCompleteLinearOrder X] [ConditionallyCompleteLinearOrder E] - [OrderTopology X] [OrderTopology E] [SecondCountableTopology E] + [OrderTopology X] [OrderTopology E] [SecondCountableTopology E] {p : ℝ≥0∞} -theorem MonotoneOn.integrableOn_of_measure_ne_top (hmono : MonotoneOn f s) {a b : X} - (ha : IsLeast s a) (hb : IsGreatest s b) (hs : μ s ≠ ∞) (h's : MeasurableSet s) : - IntegrableOn f s μ := by +theorem MonotoneOn.memℒp_top (hmono : MonotoneOn f s) {a b : X} + (ha : IsLeast s a) (hb : IsGreatest s b) (h's : MeasurableSet s) : + Memℒp f ∞ (μ.restrict s) := by borelize E - obtain rfl | _ := s.eq_empty_or_nonempty - · exact integrableOn_empty have hbelow : BddBelow (f '' s) := ⟨f a, fun x ⟨y, hy, hyx⟩ => hyx ▸ hmono ha.1 hy (ha.2 hy)⟩ have habove : BddAbove (f '' s) := ⟨f b, fun x ⟨y, hy, hyx⟩ => hyx ▸ hmono hy hb.1 (hb.2 hy)⟩ have : IsBounded (f '' s) := Metric.isBounded_of_bddAbove_of_bddBelow habove hbelow rcases isBounded_iff_forall_norm_le.mp this with ⟨C, hC⟩ - have A : IntegrableOn (fun _ => C) s μ := by - simp only [hs.lt_top, integrableOn_const, or_true] - exact - Integrable.mono' A (aemeasurable_restrict_of_monotoneOn h's hmono).aestronglyMeasurable - ((ae_restrict_iff' h's).mpr <| ae_of_all _ fun y hy => hC (f y) (mem_image_of_mem f hy)) + have A : Memℒp (fun _ => C) ⊤ (μ.restrict s) := memℒp_top_const _ + apply Memℒp.mono A (aemeasurable_restrict_of_monotoneOn h's hmono).aestronglyMeasurable + apply (ae_restrict_iff' h's).mpr + apply ae_of_all _ fun y hy ↦ ?_ + exact (hC _ (mem_image_of_mem f hy)).trans (le_abs_self _) -theorem MonotoneOn.integrableOn_isCompact [IsFiniteMeasureOnCompacts μ] (hs : IsCompact s) - (hmono : MonotoneOn f s) : IntegrableOn f s μ := by +theorem MonotoneOn.memℒp_of_measure_ne_top (hmono : MonotoneOn f s) {a b : X} + (ha : IsLeast s a) (hb : IsGreatest s b) (hs : μ s ≠ ∞) (h's : MeasurableSet s) : + Memℒp f p (μ.restrict s) := + (hmono.memℒp_top ha hb h's).memℒp_of_exponent_le_of_measure_support_ne_top (s := univ) + (by simp) (by simpa using hs) le_top + +theorem MonotoneOn.memℒp_isCompact [IsFiniteMeasureOnCompacts μ] (hs : IsCompact s) + (hmono : MonotoneOn f s) : Memℒp f p (μ.restrict s) := by obtain rfl | h := s.eq_empty_or_nonempty - · exact integrableOn_empty - · exact - hmono.integrableOn_of_measure_ne_top (hs.isLeast_sInf h) (hs.isGreatest_sSup h) - hs.measure_lt_top.ne hs.measurableSet + · simp + · exact hmono.memℒp_of_measure_ne_top (hs.isLeast_sInf h) (hs.isGreatest_sSup h) + hs.measure_lt_top.ne hs.measurableSet + +theorem AntitoneOn.memℒp_top (hanti : AntitoneOn f s) {a b : X} + (ha : IsLeast s a) (hb : IsGreatest s b) (h's : MeasurableSet s) : + Memℒp f ∞ (μ.restrict s) := + MonotoneOn.memℒp_top (E := Eᵒᵈ) hanti ha hb h's + +theorem AntitoneOn.memℒp_of_measure_ne_top (hanti : AntitoneOn f s) {a b : X} + (ha : IsLeast s a) (hb : IsGreatest s b) (hs : μ s ≠ ∞) (h's : MeasurableSet s) : + Memℒp f p (μ.restrict s) := + MonotoneOn.memℒp_of_measure_ne_top (E := Eᵒᵈ) hanti ha hb hs h's + +theorem AntitoneOn.memℒp_isCompact [IsFiniteMeasureOnCompacts μ] (hs : IsCompact s) + (hanti : AntitoneOn f s) : Memℒp f p (μ.restrict s) := + MonotoneOn.memℒp_isCompact (E := Eᵒᵈ) hs hanti + +theorem MonotoneOn.integrableOn_of_measure_ne_top (hmono : MonotoneOn f s) {a b : X} + (ha : IsLeast s a) (hb : IsGreatest s b) (hs : μ s ≠ ∞) (h's : MeasurableSet s) : + IntegrableOn f s μ := + memℒp_one_iff_integrable.1 (hmono.memℒp_of_measure_ne_top ha hb hs h's) + +theorem MonotoneOn.integrableOn_isCompact [IsFiniteMeasureOnCompacts μ] (hs : IsCompact s) + (hmono : MonotoneOn f s) : IntegrableOn f s μ := + memℒp_one_iff_integrable.1 (hmono.memℒp_isCompact hs) theorem AntitoneOn.integrableOn_of_measure_ne_top (hanti : AntitoneOn f s) {a b : X} (ha : IsLeast s a) (hb : IsGreatest s b) (hs : μ s ≠ ∞) (h's : MeasurableSet s) : IntegrableOn f s μ := - hanti.dual_right.integrableOn_of_measure_ne_top ha hb hs h's + memℒp_one_iff_integrable.1 (hanti.memℒp_of_measure_ne_top ha hb hs h's) theorem AntioneOn.integrableOn_isCompact [IsFiniteMeasureOnCompacts μ] (hs : IsCompact s) (hanti : AntitoneOn f s) : IntegrableOn f s μ := - hanti.dual_right.integrableOn_isCompact (E := Eᵒᵈ) hs + memℒp_one_iff_integrable.1 (hanti.memℒp_isCompact hs) theorem Monotone.locallyIntegrable [IsLocallyFiniteMeasure μ] (hmono : Monotone f) : LocallyIntegrable f μ := by diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index fdfe71f12dff1..bc362c375e84d 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -311,6 +311,8 @@ theorem eLpNorm_measure_zero {f : α → F} : eLpNorm f p (0 : Measure α) = 0 : @[deprecated (since := "2024-07-27")] alias snorm_measure_zero := eLpNorm_measure_zero +@[simp] lemma memℒp_measure_zero {f : α → F} : Memℒp f p (0 : Measure α) := by simp [Memℒp] + end Zero section Neg diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean index cc7d839a0520c..1df002f2d7299 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean @@ -357,4 +357,39 @@ theorem Memℒp.smul_of_top_left {p : ℝ≥0∞} {f : α → E} {φ : α → end BoundedSMul +section Mul + +variable {α : Type*} [MeasurableSpace α] {𝕜 : Type*} [NormedRing 𝕜] {μ : Measure α} + {p q r : ℝ≥0∞} {f : α → 𝕜} {φ : α → 𝕜} + +theorem Memℒp.mul (hf : Memℒp f r μ) (hφ : Memℒp φ q μ) (hpqr : 1 / p = 1 / q + 1 / r) : + Memℒp (φ * f) p μ := + Memℒp.smul hf hφ hpqr + +/-- Variant of `Memℒp.mul` where the function is written as `fun x ↦ φ x * f x` +instead of `φ * f`. -/ +theorem Memℒp.mul' (hf : Memℒp f r μ) (hφ : Memℒp φ q μ) (hpqr : 1 / p = 1 / q + 1 / r) : + Memℒp (fun x ↦ φ x * f x) p μ := + Memℒp.smul hf hφ hpqr + +theorem Memℒp.mul_of_top_right (hf : Memℒp f p μ) (hφ : Memℒp φ ∞ μ) : Memℒp (φ * f) p μ := + Memℒp.smul_of_top_right hf hφ + +/-- Variant of `Memℒp.mul_of_top_right` where the function is written as `fun x ↦ φ x * f x` +instead of `φ * f`. -/ +theorem Memℒp.mul_of_top_right' (hf : Memℒp f p μ) (hφ : Memℒp φ ∞ μ) : + Memℒp (fun x ↦ φ x * f x) p μ := + Memℒp.smul_of_top_right hf hφ + +theorem Memℒp.mul_of_top_left (hf : Memℒp f ∞ μ) (hφ : Memℒp φ p μ) : Memℒp (φ * f) p μ := + Memℒp.smul_of_top_left hf hφ + +/-- Variant of `Memℒp.mul_of_top_left` where the function is written as `fun x ↦ φ x * f x` +instead of `φ * f`. -/ +theorem Memℒp.mul_of_top_left' (hf : Memℒp f ∞ μ) (hφ : Memℒp φ p μ) : + Memℒp (fun x ↦ φ x * f x) p μ := + Memℒp.smul_of_top_left hf hφ + +end Mul + end MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index b2da38021c0f9..757c23abb972f 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -375,6 +375,21 @@ theorem of_uncurry_right [TopologicalSpace β] {_ : MeasurableSpace α} {_ : Mea StronglyMeasurable fun x => f x y := hf.comp_measurable measurable_prod_mk_right +protected theorem prod_swap {_ : MeasurableSpace α} {_ : MeasurableSpace β} [TopologicalSpace γ] + {f : β × α → γ} (hf : StronglyMeasurable f) : + StronglyMeasurable (fun z : α × β => f z.swap) := + hf.comp_measurable measurable_swap + +protected theorem fst {_ : MeasurableSpace α} [mβ : MeasurableSpace β] [TopologicalSpace γ] + {f : α → γ} (hf : StronglyMeasurable f) : + StronglyMeasurable (fun z : α × β => f z.1) := + hf.comp_measurable measurable_fst + +protected theorem snd [mα : MeasurableSpace α] {_ : MeasurableSpace β} [TopologicalSpace γ] + {f : β → γ} (hf : StronglyMeasurable f) : + StronglyMeasurable (fun z : α × β => f z.2) := + hf.comp_measurable measurable_snd + section Arithmetic variable {mα : MeasurableSpace α} [TopologicalSpace β] From 894f765e7617adfc3678b8797d806c2d899bf356 Mon Sep 17 00:00:00 2001 From: tukamilano Date: Fri, 3 Jan 2025 11:53:10 +0000 Subject: [PATCH 111/189] =?UTF-8?q?feat:=20`fun=20=CF=89=20=E2=86=A6=20exp?= =?UTF-8?q?=20(t=20*=20X=20=CF=89)`=20is=20integrable=20if=20`X`=20is=20bo?= =?UTF-8?q?unded=20above=20(#20017)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I extracted a part of the [pull request of Hoeffding's lemma](https://github.com/leanprover-community/mathlib4/pull/18091). Co-authored-by: tukamilano <107251402+tukamilano@users.noreply.github.com> --- Mathlib/MeasureTheory/Function/L1Space.lean | 11 +++++++++++ Mathlib/Probability/Moments.lean | 12 +++++++++++- 2 files changed, 22 insertions(+), 1 deletion(-) diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space.lean index ca4c26d3920da..6d78b35cdd579 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space.lean @@ -158,6 +158,12 @@ theorem hasFiniteIntegral_const [IsFiniteMeasure μ] (c : β) : HasFiniteIntegral (fun _ : α => c) μ := hasFiniteIntegral_const_iff.2 (Or.inr <| measure_lt_top _ _) +theorem HasFiniteIntegral.of_mem_Icc [IsFiniteMeasure μ] (a b : ℝ) {X : α → ℝ} + (h : ∀ᵐ ω ∂μ, X ω ∈ Set.Icc a b) : + HasFiniteIntegral X μ := by + apply (hasFiniteIntegral_const (max ‖a‖ ‖b‖)).mono' + filter_upwards [h.mono fun ω h ↦ h.1, h.mono fun ω h ↦ h.2] with ω using abs_le_max_abs_abs + theorem hasFiniteIntegral_of_bounded [IsFiniteMeasure μ] {f : α → β} {C : ℝ} (hC : ∀ᵐ a ∂μ, ‖f a‖ ≤ C) : HasFiniteIntegral f μ := (hasFiniteIntegral_const C).mono' hC @@ -450,6 +456,11 @@ theorem integrable_const_iff {c : β} : Integrable (fun _ : α => c) μ ↔ c = have : AEStronglyMeasurable (fun _ : α => c) μ := aestronglyMeasurable_const rw [Integrable, and_iff_right this, hasFiniteIntegral_const_iff] +theorem Integrable.of_mem_Icc [IsFiniteMeasure μ] (a b : ℝ) {X : α → ℝ} (hX : AEMeasurable X μ) + (h : ∀ᵐ ω ∂μ, X ω ∈ Set.Icc a b) : + Integrable X μ := + ⟨hX.aestronglyMeasurable, .of_mem_Icc a b h⟩ + @[simp] theorem integrable_const [IsFiniteMeasure μ] (c : β) : Integrable (fun _ : α => c) μ := integrable_const_iff.2 <| Or.inr <| measure_lt_top _ _ diff --git a/Mathlib/Probability/Moments.lean b/Mathlib/Probability/Moments.lean index 41fb7a9465a19..e6ac3e9af4f90 100644 --- a/Mathlib/Probability/Moments.lean +++ b/Mathlib/Probability/Moments.lean @@ -31,7 +31,6 @@ import Mathlib.Probability.Variance `ProbabilityTheory.measure_ge_le_exp_mul_mgf` and `ProbabilityTheory.measure_le_le_exp_mul_mgf` for versions of these results using `mgf` instead of `cgf`. - -/ @@ -356,4 +355,15 @@ lemma mgf_dirac {x : ℝ} (hX : μ.map X = .dirac x) (t : ℝ) : mgf X μ t = ex end MomentGeneratingFunction +lemma aemeasurable_exp_mul {X : Ω → ℝ} (t : ℝ) (hX : AEMeasurable X μ) : + AEStronglyMeasurable (fun ω ↦ rexp (t * X ω)) μ := + (measurable_exp.comp_aemeasurable (hX.const_mul t)).aestronglyMeasurable + +lemma integrable_exp_mul_of_le [IsFiniteMeasure μ] {X : Ω → ℝ} (t b : ℝ) (ht : 0 ≤ t) + (hX : AEMeasurable X μ) (hb : ∀ᵐ ω ∂μ, X ω ≤ b) : + Integrable (fun ω ↦ exp (t * X ω)) μ := by + refine .of_mem_Icc 0 (rexp (t * b)) (measurable_exp.comp_aemeasurable (hX.const_mul t)) ?_ + filter_upwards [hb] with ω hb + exact ⟨by positivity, by gcongr⟩ + end ProbabilityTheory From a8f082c87c9ace281ef3117c97668cc0d436de07 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Fri, 3 Jan 2025 11:53:11 +0000 Subject: [PATCH 112/189] chore(CategoryTheory/Shift): fix NatTrans.CommShift (#20364) This PR improves some proofs about `NatTrans.CommShift` (the commutation of a natural transformation to shifts). The field `comm'` of `NatTrans.CommShift` had an unnecessary `'` in it, it is renamed `shift_comm`, and API lemmas are moved from `NatTrans.CommShift` to the `NatTrans` namespace (where `comm` becomes `shift_comm` and `comm_app` becomes `shift_app_comm`). --- Mathlib/CategoryTheory/Shift/Adjunction.lean | 6 +- Mathlib/CategoryTheory/Shift/CommShift.lean | 70 +++++++++---------- Mathlib/CategoryTheory/Shift/Quotient.lean | 2 +- .../CategoryTheory/Shift/SingleFunctors.lean | 2 +- .../CategoryTheory/Triangulated/Functor.lean | 2 +- 5 files changed, 39 insertions(+), 43 deletions(-) diff --git a/Mathlib/CategoryTheory/Shift/Adjunction.lean b/Mathlib/CategoryTheory/Shift/Adjunction.lean index ec22e3951bc7b..f893e59286f66 100644 --- a/Mathlib/CategoryTheory/Shift/Adjunction.lean +++ b/Mathlib/CategoryTheory/Shift/Adjunction.lean @@ -215,7 +215,7 @@ lemma mk' (h : NatTrans.CommShift adj.unit A) : refine (compatibilityCounit_of_compatibilityUnit adj _ _ (fun X ↦ ?_) _).symm simpa only [NatTrans.comp_app, Functor.commShiftIso_id_hom_app, whiskerRight_app, id_comp, - Functor.commShiftIso_comp_hom_app] using congr_app (h.comm' a) X⟩ + Functor.commShiftIso_comp_hom_app] using congr_app (h.shift_comm a) X⟩ end CommShift @@ -227,14 +227,14 @@ lemma shift_unit_app [adj.CommShift A] (a : A) (X : C) : adj.unit.app (X⟦a⟧) ≫ G.map ((F.commShiftIso a).hom.app X) ≫ (G.commShiftIso a).hom.app (F.obj X) := by - simpa [Functor.commShiftIso_comp_hom_app] using NatTrans.CommShift.comm_app adj.unit a X + simpa [Functor.commShiftIso_comp_hom_app] using NatTrans.shift_app_comm adj.unit a X @[reassoc] lemma shift_counit_app [adj.CommShift A] (a : A) (Y : D) : (adj.counit.app Y)⟦a⟧' = (F.commShiftIso a).inv.app (G.obj Y) ≫ F.map ((G.commShiftIso a).inv.app Y) ≫ adj.counit.app (Y⟦a⟧) := by - have eq := NatTrans.CommShift.comm_app adj.counit a Y + have eq := NatTrans.shift_app_comm adj.counit a Y simp only [Functor.comp_obj, Functor.id_obj, Functor.commShiftIso_comp_hom_app, assoc, Functor.commShiftIso_id_hom_app, comp_id] at eq simp only [← eq, Functor.comp_obj, Functor.id_obj, ← F.map_comp_assoc, Iso.inv_hom_id_app, diff --git a/Mathlib/CategoryTheory/Shift/CommShift.lean b/Mathlib/CategoryTheory/Shift/CommShift.lean index 6aabe3d9c6729..fb958733c0834 100644 --- a/Mathlib/CategoryTheory/Shift/CommShift.lean +++ b/Mathlib/CategoryTheory/Shift/CommShift.lean @@ -264,90 +264,86 @@ variable {C D E J : Type*} [Category C] [Category D] [Category E] [Category J] which commute with a shift by an additive monoid `A`, this typeclass asserts a compatibility of `τ` with these shifts. -/ class CommShift : Prop where - comm' (a : A) : (F₁.commShiftIso a).hom ≫ whiskerRight τ _ = - whiskerLeft _ τ ≫ (F₂.commShiftIso a).hom - -namespace CommShift + shift_comm (a : A) : (F₁.commShiftIso a).hom ≫ whiskerRight τ _ = + whiskerLeft _ τ ≫ (F₂.commShiftIso a).hom := by aesop_cat section -variable {A} -variable [NatTrans.CommShift τ A] +variable {A} [NatTrans.CommShift τ A] -lemma comm (a : A) : (F₁.commShiftIso a).hom ≫ whiskerRight τ _ = - whiskerLeft _ τ ≫ (F₂.commShiftIso a).hom := by - apply comm' +@[reassoc] +lemma shift_comm (a : A) : + (F₁.commShiftIso a).hom ≫ whiskerRight τ _ = + whiskerLeft _ τ ≫ (F₂.commShiftIso a).hom := by + apply CommShift.shift_comm @[reassoc] -lemma comm_app (a : A) (X : C) : +lemma shift_app_comm (a : A) (X : C) : (F₁.commShiftIso a).hom.app X ≫ (τ.app X)⟦a⟧' = τ.app (X⟦a⟧) ≫ (F₂.commShiftIso a).hom.app X := - NatTrans.congr_app (comm τ a) X + congr_app (shift_comm τ a) X @[reassoc] lemma shift_app (a : A) (X : C) : (τ.app X)⟦a⟧' = (F₁.commShiftIso a).inv.app X ≫ τ.app (X⟦a⟧) ≫ (F₂.commShiftIso a).hom.app X := by - rw [← comm_app, Iso.inv_hom_id_app_assoc] + rw [← shift_app_comm, Iso.inv_hom_id_app_assoc] @[reassoc] lemma app_shift (a : A) (X : C) : τ.app (X⟦a⟧) = (F₁.commShiftIso a).hom.app X ≫ (τ.app X)⟦a⟧' ≫ (F₂.commShiftIso a).inv.app X := by - erw [comm_app_assoc, Iso.hom_inv_id_app, Category.comp_id] + simp [shift_app_comm_assoc τ a X] + +@[deprecated (since := "2024-12-31")] alias CommShift.comm' := shift_comm +@[deprecated (since := "2024-12-31")] alias CommShift.comm := shift_comm +@[deprecated (since := "2024-12-31")] alias CommShift.comm_app := shift_app_comm +@[deprecated (since := "2024-12-31")] alias CommShift.shift_app := shift_app +@[deprecated (since := "2024-12-31")] alias CommShift.app_shift := app_shift end +namespace CommShift + instance of_iso_inv [NatTrans.CommShift e.hom A] : NatTrans.CommShift e.inv A := ⟨fun a => by ext X dsimp - rw [← cancel_epi (e.hom.app (X⟦a⟧)), e.hom_inv_id_app_assoc, ← comm_app_assoc, - ← Functor.map_comp, e.hom_inv_id_app, Functor.map_id] - rw [Category.comp_id]⟩ + rw [← cancel_epi (e.hom.app (X⟦a⟧)), e.hom_inv_id_app_assoc, ← shift_app_comm_assoc, + ← Functor.map_comp, e.hom_inv_id_app, Functor.map_id, Category.comp_id]⟩ lemma of_isIso [IsIso τ] [NatTrans.CommShift τ A] : NatTrans.CommShift (inv τ) A := by - haveI : NatTrans.CommShift (asIso τ).hom A := by - dsimp - infer_instance + haveI : NatTrans.CommShift (asIso τ).hom A := by assumption change NatTrans.CommShift (asIso τ).inv A infer_instance variable (F₁) in -instance id : NatTrans.CommShift (𝟙 F₁) A := ⟨by aesop_cat⟩ +instance id : NatTrans.CommShift (𝟙 F₁) A where + +attribute [local simp] Functor.commShiftIso_comp_hom_app + shift_app_comm shift_app_comm_assoc instance comp [NatTrans.CommShift τ A] [NatTrans.CommShift τ' A] : - NatTrans.CommShift (τ ≫ τ') A := ⟨fun a => by - ext X - simp [comm_app_assoc, comm_app]⟩ + NatTrans.CommShift (τ ≫ τ') A where instance whiskerRight [NatTrans.CommShift τ A] : NatTrans.CommShift (whiskerRight τ G) A := ⟨fun a => by ext X - simp only [Functor.comp_obj, whiskerRight_twice, comp_app, + simp only [whiskerRight_twice, comp_app, whiskerRight_app, Functor.comp_map, whiskerLeft_app, - Functor.commShiftIso_comp_hom_app, Category.assoc] - erw [← NatTrans.naturality] - dsimp - simp only [← G.map_comp_assoc, comm_app]⟩ - -variable {G G'} (F₁) + Functor.commShiftIso_comp_hom_app, Category.assoc, + ← Functor.commShiftIso_hom_naturality, + ← G.map_comp_assoc, shift_app_comm]⟩ instance whiskerLeft [NatTrans.CommShift τ'' A] : - NatTrans.CommShift (whiskerLeft F₁ τ'') A := ⟨fun a => by - ext X - simp only [Functor.comp_obj, comp_app, whiskerRight_app, whiskerLeft_app, whiskerLeft_twice, - Functor.commShiftIso_comp_hom_app, Category.assoc, ← NatTrans.naturality_assoc, comm_app]⟩ + NatTrans.CommShift (whiskerLeft F₁ τ'') A where instance associator : CommShift (Functor.associator F₁ G H).hom A where - comm' a := by ext X; simp [Functor.commShiftIso_comp_hom_app] instance leftUnitor : CommShift F₁.leftUnitor.hom A where - comm' a := by ext X; simp [Functor.commShiftIso_comp_hom_app] instance rightUnitor : CommShift F₁.rightUnitor.hom A where - comm' a := by ext X; simp [Functor.commShiftIso_comp_hom_app] end CommShift diff --git a/Mathlib/CategoryTheory/Shift/Quotient.lean b/Mathlib/CategoryTheory/Shift/Quotient.lean index f789f82c2883c..3ea668302cc4c 100644 --- a/Mathlib/CategoryTheory/Shift/Quotient.lean +++ b/Mathlib/CategoryTheory/Shift/Quotient.lean @@ -136,7 +136,7 @@ noncomputable instance liftCommShift : instance liftCommShift_compatibility : NatTrans.CommShift (Quotient.lift.isLift r F hF).hom A where - comm' a := by + shift_comm a := by ext X dsimp erw [Functor.map_id, id_comp, comp_id] diff --git a/Mathlib/CategoryTheory/Shift/SingleFunctors.lean b/Mathlib/CategoryTheory/Shift/SingleFunctors.lean index 115e77228da25..bb542f2bd9479 100644 --- a/Mathlib/CategoryTheory/Shift/SingleFunctors.lean +++ b/Mathlib/CategoryTheory/Shift/SingleFunctors.lean @@ -250,7 +250,7 @@ def postcompIsoOfIso {G G' : D ⥤ E} (e : G ≅ G') [G.CommShift A] [G'.CommShi isoMk (fun a => isoWhiskerLeft (F.functor a) e) (fun n a a' ha' => by ext X dsimp - simp [NatTrans.CommShift.shift_app e.hom n]) + simp [NatTrans.shift_app e.hom n]) end SingleFunctors diff --git a/Mathlib/CategoryTheory/Triangulated/Functor.lean b/Mathlib/CategoryTheory/Triangulated/Functor.lean index 2bafc57433276..c7c860c35beec 100644 --- a/Mathlib/CategoryTheory/Triangulated/Functor.lean +++ b/Mathlib/CategoryTheory/Triangulated/Functor.lean @@ -135,7 +135,7 @@ def mapTriangleIso {F₁ F₂ : C ⥤ D} (e : F₁ ≅ F₂) [F₁.CommShift ℤ NatIso.ofComponents (fun T => Triangle.isoMk _ _ (e.app _) (e.app _) (e.app _) (by simp) (by simp) (by dsimp - simp only [assoc, NatTrans.CommShift.comm_app e.hom (1 : ℤ) T.obj₁, + simp only [assoc, NatTrans.shift_app_comm e.hom (1 : ℤ) T.obj₁, NatTrans.naturality_assoc])) (by aesop_cat) end Additive From 8f95bde32302ac3cb18e1d2d1318912291fca02c Mon Sep 17 00:00:00 2001 From: Quang Dao Date: Fri, 3 Jan 2025 12:20:40 +0000 Subject: [PATCH 113/189] feat(Order/Interval/Finset): Lemmas for intervals with bot/top (#20360) Add some lemmas on intervals containing bot / top elements, mirroring existing results in `Order/Interval/Set/Basic.lean`. Co-authored-by: Quang Dao --- Mathlib/Order/Interval/Finset/Basic.lean | 48 ++++++++++++++++++++++++ 1 file changed, 48 insertions(+) diff --git a/Mathlib/Order/Interval/Finset/Basic.lean b/Mathlib/Order/Interval/Finset/Basic.lean index 43829446d42e4..ef23dafd228db 100644 --- a/Mathlib/Order/Interval/Finset/Basic.lean +++ b/Mathlib/Order/Interval/Finset/Basic.lean @@ -308,6 +308,17 @@ section LocallyFiniteOrderTop variable [LocallyFiniteOrderTop α] +@[simp] +theorem Ioi_eq_empty : Ioi a = ∅ ↔ IsMax a := by + rw [← coe_eq_empty, coe_Ioi, Set.Ioi_eq_empty_iff] + +@[simp] +theorem Ioi_top [OrderTop α] : Ioi (⊤ : α) = ∅ := Ioi_eq_empty.mpr isMax_top + +@[simp] +theorem Ici_bot [OrderBot α] [Fintype α] : Ici (⊥ : α) = univ := by + ext a; simp only [mem_Ici, bot_le, mem_univ] + @[simp, aesop safe apply (rule_sets := [finsetNonempty])] lemma nonempty_Ici : (Ici a).Nonempty := ⟨a, mem_Ici.2 le_rfl⟩ @[simp] @@ -348,6 +359,16 @@ section LocallyFiniteOrderBot variable [LocallyFiniteOrderBot α] +@[simp] +theorem Iio_eq_empty : Iio a = ∅ ↔ IsMin a := Ioi_eq_empty (α := αᵒᵈ) + +@[simp] +theorem Iio_bot [OrderBot α] : Iio (⊥ : α) = ∅ := Iio_eq_empty.mpr isMin_bot + +@[simp] +theorem Iic_top [OrderTop α] [Fintype α] : Iic (⊤ : α) = univ := by + ext a; simp only [mem_Iic, le_top, mem_univ] + @[simp, aesop safe apply (rule_sets := [finsetNonempty])] lemma nonempty_Iic : (Iic a).Nonempty := ⟨a, mem_Iic.2 le_rfl⟩ @[simp] @@ -425,6 +446,27 @@ theorem filter_ge_eq_Iic [DecidablePred (· ≤ a)] : ({x | x ≤ a} : Finset _) end LocallyFiniteOrderBot +section LocallyFiniteOrder + +variable [LocallyFiniteOrder α] + +@[simp] +theorem Icc_bot [OrderBot α] : Icc (⊥ : α) a = Iic a := rfl + +@[simp] +theorem Icc_top [OrderTop α] : Icc a (⊤ : α) = Ici a := rfl + +@[simp] +theorem Ico_bot [OrderBot α] : Ico (⊥ : α) a = Iio a := rfl + +@[simp] +theorem Ioc_top [OrderTop α] : Ioc a (⊤ : α) = Ioi a := rfl + +theorem Icc_bot_top [BoundedOrder α] [Fintype α] : Icc (⊥ : α) (⊤ : α) = univ := by + rw [Icc_bot, Iic_top] + +end LocallyFiniteOrder + variable [LocallyFiniteOrderTop α] [LocallyFiniteOrderBot α] theorem disjoint_Ioi_Iio (a : α) : Disjoint (Ioi a) (Iio a) := @@ -446,6 +488,12 @@ theorem Icc_eq_singleton_iff : Icc a b = {c} ↔ a = c ∧ b = c := by theorem Ico_disjoint_Ico_consecutive (a b c : α) : Disjoint (Ico a b) (Ico b c) := disjoint_left.2 fun _ hab hbc => (mem_Ico.mp hab).2.not_le (mem_Ico.mp hbc).1 +@[simp] +theorem Ici_top [OrderTop α] : Ici (⊤ : α) = {⊤} := Icc_eq_singleton_iff.2 ⟨rfl, rfl⟩ + +@[simp] +theorem Iic_bot [OrderBot α] : Iic (⊥ : α) = {⊥} := Icc_eq_singleton_iff.2 ⟨rfl, rfl⟩ + section DecidableEq variable [DecidableEq α] From 37515f7950787e26399ad11e90f1b3e50b2c9139 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Fri, 3 Jan 2025 12:52:16 +0000 Subject: [PATCH 114/189] feat(Algebra/Group): Units in pi types (#20279) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Algebra/Group/Pi/Units.lean | 39 +++++++++++++++++++++++++++++ 2 files changed, 40 insertions(+) create mode 100644 Mathlib/Algebra/Group/Pi/Units.lean diff --git a/Mathlib.lean b/Mathlib.lean index fb780d7561c21..baddcd1b0a810 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -305,6 +305,7 @@ import Mathlib.Algebra.Group.Opposite import Mathlib.Algebra.Group.PNatPowAssoc import Mathlib.Algebra.Group.Pi.Basic import Mathlib.Algebra.Group.Pi.Lemmas +import Mathlib.Algebra.Group.Pi.Units import Mathlib.Algebra.Group.Pointwise.Finset.Basic import Mathlib.Algebra.Group.Pointwise.Finset.Interval import Mathlib.Algebra.Group.Pointwise.Set.Basic diff --git a/Mathlib/Algebra/Group/Pi/Units.lean b/Mathlib/Algebra/Group/Pi/Units.lean new file mode 100644 index 0000000000000..ae0ecae78a6f7 --- /dev/null +++ b/Mathlib/Algebra/Group/Pi/Units.lean @@ -0,0 +1,39 @@ +/- +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.Algebra.Group.Pi.Basic +import Mathlib.Algebra.Group.Units.Defs +import Mathlib.Algebra.Group.Equiv.Basic +import Mathlib.Util.Delaborators + +/-! # Units in pi types -/ + +variable {ι : Type*} {M : ι → Type*} [∀ i, Monoid (M i)] {x : Π i, M i} + +open Units in +/-- The monoid equivalence between units of a product, +and the product of the units of each monoid. -/ +@[to_additive (attr := simps) + "The additive-monoid equivalence between (additive) units of a product, + and the product of the (additive) units of each monoid."] +def MulEquiv.piUnits : (Π i, M i)ˣ ≃* Π i, (M i)ˣ where + toFun f i := ⟨f.val i, f.inv i, congr_fun f.val_inv i, congr_fun f.inv_val i⟩ + invFun f := ⟨(val <| f ·), (inv <| f ·), funext (val_inv <| f ·), funext (inv_val <| f ·)⟩ + left_inv _ := rfl + right_inv _ := rfl + map_mul' _ _ := rfl + +@[to_additive] +lemma Pi.isUnit_iff : + IsUnit x ↔ ∀ i, IsUnit (x i) := by + simp_rw [isUnit_iff_exists, funext_iff, ← forall_and] + exact Classical.skolem (p := fun i y ↦ x i * y = 1 ∧ y * x i = 1).symm + +@[to_additive] +alias ⟨IsUnit.apply, _⟩ := Pi.isUnit_iff + +@[to_additive] +lemma IsUnit.val_inv_apply (hx : IsUnit x) (i : ι) : (hx.unit⁻¹).1 i = (hx.apply i).unit⁻¹ := by + rw [← Units.inv_eq_val_inv, ← MulEquiv.val_inv_piUnits_apply]; congr; ext; rfl From 860c303880643ae1bf0b5042620a7813b50613c6 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Fri, 3 Jan 2025 12:52:17 +0000 Subject: [PATCH 115/189] feat: add `Prod.isUnit_iff` (#20432) --- Mathlib/Algebra/Group/Prod.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Algebra/Group/Prod.lean b/Mathlib/Algebra/Group/Prod.lean index e6e9519ae6483..739752b4ea517 100644 --- a/Mathlib/Algebra/Group/Prod.lean +++ b/Mathlib/Algebra/Group/Prod.lean @@ -671,6 +671,11 @@ def prodUnits : (M × N)ˣ ≃* Mˣ × Nˣ where exact ⟨rfl, rfl⟩ map_mul' := MonoidHom.map_mul _ +@[to_additive] +lemma _root_.Prod.isUnit_iff {x : M × N} : IsUnit x ↔ IsUnit x.1 ∧ IsUnit x.2 where + mp h := ⟨(prodUnits h.unit).1.isUnit, (prodUnits h.unit).2.isUnit⟩ + mpr h := (prodUnits.symm (h.1.unit, h.2.unit)).isUnit + end end MulEquiv From e3a43bba8a3ac9036f2f88318bb13ddcc6b73ba3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 3 Jan 2025 13:31:10 +0000 Subject: [PATCH 116/189] =?UTF-8?q?chore:=20fix=20the=20`Integrable[m?= =?UTF-8?q?=CE=B1]`=20notation=20(#20434)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit It was broken in #20122. From GibbsMeasure --- Mathlib/MeasureTheory/Function/L1Space.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space.lean index 6d78b35cdd579..30d728e7e06fc 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space.lean @@ -413,7 +413,7 @@ def Integrable {α} {_ : MeasurableSpace α} (f : α → ε) AEStronglyMeasurable f μ ∧ HasFiniteIntegral f μ /-- Notation for `Integrable` with respect to a non-standard σ-algebra. -/ -scoped notation "Integrable[" mα "]" => @Integrable _ _ _ mα +scoped notation "Integrable[" mα "]" => @Integrable _ _ _ _ mα theorem memℒp_one_iff_integrable {f : α → β} : Memℒp f 1 μ ↔ Integrable f μ := by simp_rw [Integrable, hasFiniteIntegral_iff_nnnorm, Memℒp, eLpNorm_one_eq_lintegral_nnnorm] From aba55dcb15076c646273f4459f055fd04d4f3fdf Mon Sep 17 00:00:00 2001 From: Fabrizio Barroero Date: Fri, 3 Jan 2025 13:44:58 +0000 Subject: [PATCH 117/189] feat: some results about roots of linear polynomials (#20386) We prove some results for roots of linear polynomials. --- Mathlib/Algebra/Polynomial/FieldDivision.lean | 21 ++++++++++++------- Mathlib/Algebra/Polynomial/Roots.lean | 15 +++++++++++++ 2 files changed, 29 insertions(+), 7 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/FieldDivision.lean b/Mathlib/Algebra/Polynomial/FieldDivision.lean index e5294cbb83106..36d71e8b2b9ee 100644 --- a/Mathlib/Algebra/Polynomial/FieldDivision.lean +++ b/Mathlib/Algebra/Polynomial/FieldDivision.lean @@ -496,14 +496,21 @@ theorem rootSet_prod [CommRing S] [IsDomain S] [Algebra R S] {ι : Type*} (f : rw [Polynomial.map_prod, roots_prod, Finset.bind_toFinset, s.val_toFinset, Finset.coe_biUnion] rwa [← Polynomial.map_prod, Ne, Polynomial.map_eq_zero] +theorem roots_C_mul_X_sub_C (b : R) (ha : a ≠ 0) : (C a * X - C b).roots = {a⁻¹ * b} := by + simp [roots_C_mul_X_sub_C_of_IsUnit b ⟨a, a⁻¹, mul_inv_cancel₀ ha, inv_mul_cancel₀ ha⟩] + +theorem roots_C_mul_X_add_C (b : R) (ha : a ≠ 0) : (C a * X + C b).roots = {-(a⁻¹ * b)} := by + simp [roots_C_mul_X_add_C_of_IsUnit b ⟨a, a⁻¹, mul_inv_cancel₀ ha, inv_mul_cancel₀ ha⟩] + +theorem roots_degree_eq_one (h : degree p = 1) : p.roots = {-((p.coeff 1)⁻¹ * p.coeff 0)} := by + rw [eq_X_add_C_of_degree_le_one (show degree p ≤ 1 by rw [h])] + have : p.coeff 1 ≠ 0 := coeff_ne_zero_of_eq_degree h + simp [roots_C_mul_X_add_C _ this] + theorem exists_root_of_degree_eq_one (h : degree p = 1) : ∃ x, IsRoot p x := - ⟨-(p.coeff 0 / p.coeff 1), by - have : p.coeff 1 ≠ 0 := by - have h' := natDegree_eq_of_degree_eq_some h - change natDegree p = 1 at h'; rw [← h'] - exact mt leadingCoeff_eq_zero.1 fun h0 => by simp [h0] at h - conv in p => rw [eq_X_add_C_of_degree_le_one (show degree p ≤ 1 by rw [h])] - simp [IsRoot, mul_div_cancel₀ _ this]⟩ + ⟨-((p.coeff 1)⁻¹ * p.coeff 0), by + rw [← mem_roots (by simp [← zero_le_degree_iff, h])] + simp [roots_degree_eq_one h]⟩ theorem coeff_inv_units (u : R[X]ˣ) (n : ℕ) : ((↑u : R[X]).coeff n)⁻¹ = (↑u⁻¹ : R[X]).coeff n := by rw [eq_C_of_degree_eq_zero (degree_coe_units u), eq_C_of_degree_eq_zero (degree_coe_units u⁻¹), diff --git a/Mathlib/Algebra/Polynomial/Roots.lean b/Mathlib/Algebra/Polynomial/Roots.lean index 12df720582d0e..3a5f5ab6b5d84 100644 --- a/Mathlib/Algebra/Polynomial/Roots.lean +++ b/Mathlib/Algebra/Polynomial/Roots.lean @@ -166,6 +166,9 @@ theorem roots_X_sub_C (r : R) : roots (X - C r) = {r} := by ext s rw [count_roots, rootMultiplicity_X_sub_C, count_singleton] +@[simp] +theorem roots_X_add_C (r : R) : roots (X + C r) = {-r} := by simpa using roots_X_sub_C (-r) + @[simp] theorem roots_X : roots (X : R[X]) = {0} := by rw [← roots_X_sub_C, C_0, sub_zero] @@ -195,6 +198,18 @@ theorem roots_smul_nonzero (p : R[X]) (ha : a ≠ 0) : (a • p).roots = p.roots lemma roots_neg (p : R[X]) : (-p).roots = p.roots := by rw [← neg_one_smul R p, roots_smul_nonzero p (neg_ne_zero.mpr one_ne_zero)] +@[simp] +theorem roots_C_mul_X_sub_C_of_IsUnit (b : R) (a : Rˣ) : (C (a : R) * X - C b).roots = + {a⁻¹ * b} := by + rw [← roots_C_mul _ (Units.ne_zero a⁻¹), mul_sub, ← mul_assoc, ← C_mul, ← C_mul, + Units.inv_mul, C_1, one_mul] + exact roots_X_sub_C (a⁻¹ * b) + +@[simp] +theorem roots_C_mul_X_add_C_of_IsUnit (b : R) (a : Rˣ) : (C (a : R) * X + C b).roots = + {-(a⁻¹ * b)} := by + rw [← sub_neg_eq_add, ← C_neg, roots_C_mul_X_sub_C_of_IsUnit, mul_neg] + theorem roots_list_prod (L : List R[X]) : (0 : R[X]) ∉ L → L.prod.roots = (L : Multiset R[X]).bind roots := List.recOn L (fun _ => roots_one) fun hd tl ih H => by From 3e65a9f725fecaac2254d96270fb1fae9da13bbb Mon Sep 17 00:00:00 2001 From: Seewoo Lee Date: Fri, 3 Jan 2025 14:02:02 +0000 Subject: [PATCH 118/189] feat: Two lemmas on divisibility and coprimality of `expand` (#20143) Co-authored-by: Seewoo Lee <49933279+seewoo5@users.noreply.github.com> Co-authored-by: Junyan Xu Co-authored-by: Jineon Baek <34874130+jcpaik@users.noreply.github.com> --- Mathlib/Algebra/Polynomial/Expand.lean | 28 ++++++++++++++++++++++++++ 1 file changed, 28 insertions(+) diff --git a/Mathlib/Algebra/Polynomial/Expand.lean b/Mathlib/Algebra/Polynomial/Expand.lean index 93bd6ad7c4f68..ddcccccd030ed 100644 --- a/Mathlib/Algebra/Polynomial/Expand.lean +++ b/Mathlib/Algebra/Polynomial/Expand.lean @@ -202,6 +202,34 @@ theorem contract_expand {f : R[X]} (hp : p ≠ 0) : contract p (expand R p f) = theorem contract_one {f : R[X]} : contract 1 f = f := ext fun n ↦ by rw [coeff_contract one_ne_zero, mul_one] +@[simp] theorem contract_C (r : R) : contract p (C r) = C r := by simp [contract] + +theorem contract_add {p : ℕ} (hp : p ≠ 0) (f g : R[X]) : + contract p (f + g) = contract p f + contract p g := by + ext; simp_rw [coeff_add, coeff_contract hp, coeff_add] + +theorem contract_mul_expand {p : ℕ} (hp : p ≠ 0) (f g : R[X]) : + contract p (f * expand R p g) = contract p f * g := by + ext n + rw [coeff_contract hp, coeff_mul, coeff_mul, ← sum_subset + (s₁ := (antidiagonal n).image fun x ↦ (x.1 * p, x.2 * p)), sum_image] + · simp_rw [coeff_expand_mul hp.bot_lt, coeff_contract hp] + · intro x hx y hy eq; simpa only [Prod.ext_iff, Nat.mul_right_cancel_iff hp.bot_lt] using eq + · simp_rw [subset_iff, mem_image, mem_antidiagonal]; rintro _ ⟨x, rfl, rfl⟩; simp_rw [add_mul] + simp_rw [mem_image, mem_antidiagonal] + intro ⟨x, y⟩ eq nex + by_cases h : p ∣ y + · obtain ⟨x, rfl⟩ : p ∣ x := (Nat.dvd_add_iff_left h).mpr (eq ▸ dvd_mul_left p n) + obtain ⟨y, rfl⟩ := h + refine (nex ⟨⟨x, y⟩, (Nat.mul_right_cancel_iff hp.bot_lt).mp ?_, by simp_rw [mul_comm]⟩).elim + rw [← eq, mul_comm, mul_add] + · rw [coeff_expand hp.bot_lt, if_neg h, mul_zero] + +@[simp] theorem isCoprime_expand {f g : R[X]} {p : ℕ} (hp : p ≠ 0) : + IsCoprime (expand R p f) (expand R p g) ↔ IsCoprime f g := + ⟨fun ⟨a, b, eq⟩ ↦ ⟨contract p a, contract p b, by + simp_rw [← contract_mul_expand hp, ← contract_add hp, eq, ← C_1, contract_C]⟩, (·.map _)⟩ + section ExpChar theorem expand_contract [CharP R p] [NoZeroDivisors R] {f : R[X]} (hf : Polynomial.derivative f = 0) From 478f8eac2e65b90ddec7395d04262c8e48fa5bbf Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Fri, 3 Jan 2025 14:42:58 +0000 Subject: [PATCH 119/189] chore: use unsigned measures for Lebesgue decomposition (#20400) Signed measures require around 500 more imports than unsigned measures, because they use Bochner integration. As there is no particular reason to use signed measures to prove the Lebesgue decomposition theorem, we use unsigned measures instead. This PR does not have a big impact on imports, since the file also contains other lemmas about integrability. The effect will be seen in a future PR splitting those out. --- .../MeasureTheory/Decomposition/Lebesgue.lean | 113 ++++++++---------- .../Decomposition/UnsignedHahn.lean | 4 +- 2 files changed, 51 insertions(+), 66 deletions(-) diff --git a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean index 86fb861a6a9cb..f44f641788f35 100644 --- a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean @@ -3,9 +3,9 @@ Copyright (c) 2021 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ -import Mathlib.MeasureTheory.Measure.Sub -import Mathlib.MeasureTheory.Decomposition.SignedHahn +import Mathlib.MeasureTheory.Decomposition.UnsignedHahn import Mathlib.MeasureTheory.Function.AEEqOfIntegral +import Mathlib.MeasureTheory.Measure.Sub /-! # Lebesgue decomposition @@ -672,36 +672,34 @@ lemma rnDeriv_add (ν₁ ν₂ μ : Measure α) [IsFiniteMeasure ν₁] [IsFinit · exact ((measurable_rnDeriv _ _).add (measurable_rnDeriv _ _)).aemeasurable · exact (lintegral_rnDeriv_lt_top (ν₁ + ν₂) μ).ne -open VectorMeasure SignedMeasure - /-- If two finite measures `μ` and `ν` are not mutually singular, there exists some `ε > 0` and a measurable set `E`, such that `ν(E) > 0` and `E` is positive with respect to `μ - εν`. This lemma is useful for the Lebesgue decomposition theorem. -/ theorem exists_positive_of_not_mutuallySingular (μ ν : Measure α) [IsFiniteMeasure μ] - [IsFiniteMeasure ν] (h : ¬μ ⟂ₘ ν) : + [IsFiniteMeasure ν] (h : ¬ μ ⟂ₘ ν) : ∃ ε : ℝ≥0, 0 < ε ∧ - ∃ E : Set α, - MeasurableSet E ∧ 0 < ν E ∧ 0 ≤[E] μ.toSignedMeasure - (ε • ν).toSignedMeasure := by + ∃ E : Set α, MeasurableSet E ∧ 0 < ν E + ∧ ∀ A, MeasurableSet A → ε * ν (A ∩ E) ≤ μ (A ∩ E) := by -- for all `n : ℕ`, obtain the Hahn decomposition for `μ - (1 / n) ν` - have : - ∀ n : ℕ, ∃ i : Set α, - MeasurableSet i ∧ - 0 ≤[i] μ.toSignedMeasure - ((1 / (n + 1) : ℝ≥0) • ν).toSignedMeasure ∧ - μ.toSignedMeasure - ((1 / (n + 1) : ℝ≥0) • ν).toSignedMeasure ≤[iᶜ] 0 := by - intro; exact exists_compl_positive_negative _ - choose f hf₁ hf₂ hf₃ using this + have h_decomp (n : ℕ) : ∃ s : Set α, MeasurableSet s + ∧ (∀ t, MeasurableSet t → ((1 / (n + 1) : ℝ≥0) • ν) (t ∩ s) ≤ μ (t ∩ s)) + ∧ (∀ t, MeasurableSet t → μ (t ∩ sᶜ) ≤ ((1 / (n + 1) : ℝ≥0) • ν) (t ∩ sᶜ)) := by + obtain ⟨s, hs, hs_le, hs_ge⟩ := hahn_decomposition μ ((1 / (n + 1) : ℝ≥0) • ν) + refine ⟨s, hs, fun t ht ↦ ?_, fun t ht ↦ ?_⟩ + · exact hs_le (t ∩ s) (ht.inter hs) inter_subset_right + · exact hs_ge (t ∩ sᶜ) (ht.inter hs.compl) inter_subset_right + choose f hf₁ hf₂ hf₃ using h_decomp -- set `A` to be the intersection of all the negative parts of obtained Hahn decompositions -- and we show that `μ A = 0` let A := ⋂ n, (f n)ᶜ have hAmeas : MeasurableSet A := MeasurableSet.iInter fun n ↦ (hf₁ n).compl - have hA₂ : ∀ n : ℕ, μ.toSignedMeasure - ((1 / (n + 1) : ℝ≥0) • ν).toSignedMeasure ≤[A] 0 := by - intro n; exact restrict_le_restrict_subset _ _ (hf₁ n).compl (hf₃ n) (iInter_subset _ _) - have hA₃ : ∀ n : ℕ, μ A ≤ (1 / (n + 1) : ℝ≥0) * ν A := by - intro n - have := nonpos_of_restrict_le_zero _ (hA₂ n) - rwa [toSignedMeasure_sub_apply hAmeas, sub_nonpos, ENNReal.toReal_le_toReal] at this - exacts [measure_ne_top _ _, measure_ne_top _ _] + have hA₂ (n : ℕ) (t : Set α) (ht : MeasurableSet t) : + μ (t ∩ A) ≤ ((1 / (n + 1) : ℝ≥0) • ν) (t ∩ A) := by + specialize hf₃ n (t ∩ A) (ht.inter hAmeas) + have : A ∩ (f n)ᶜ = A := inter_eq_left.mpr (iInter_subset _ n) + rwa [inter_assoc, this] at hf₃ + have hA₃ (n : ℕ) : μ A ≤ (1 / (n + 1) : ℝ≥0) * ν A := by simpa using hA₂ n univ .univ have hμ : μ A = 0 := by lift μ A to ℝ≥0 using measure_ne_top _ _ with μA lift ν A to ℝ≥0 using measure_ne_top _ _ with νA @@ -828,8 +826,8 @@ with respect to `ν` and `μ = ξ + ν.withDensity f`. This is not an instance since this is also shown for the more general σ-finite measures with `MeasureTheory.Measure.haveLebesgueDecomposition_of_sigmaFinite`. -/ theorem haveLebesgueDecomposition_of_finiteMeasure [IsFiniteMeasure μ] [IsFiniteMeasure ν] : - HaveLebesgueDecomposition μ ν := - ⟨by + HaveLebesgueDecomposition μ ν where + lebesgue_decomposition := by have h := @exists_seq_tendsto_sSup _ _ _ _ _ (measurableLEEval ν μ) ⟨0, 0, zero_mem_measurableLE, by simp⟩ (OrderTop.bddAbove _) choose g _ hg₂ f hf₁ hf₂ using h @@ -837,17 +835,16 @@ theorem haveLebesgueDecomposition_of_finiteMeasure [IsFiniteMeasure μ] [IsFinit set ξ := ⨆ (n) (k) (_ : k ≤ n), f k with hξ -- we see that `ξ` has the largest integral among all functions in `measurableLE` have hξ₁ : sSup (measurableLEEval ν μ) = ∫⁻ a, ξ a ∂ν := by - have := - @lintegral_tendsto_of_tendsto_of_monotone _ _ ν (fun n ↦ ⨆ (k) (_ : k ≤ n), f k) + have := @lintegral_tendsto_of_tendsto_of_monotone _ _ ν (fun n ↦ ⨆ (k) (_ : k ≤ n), f k) (⨆ (n) (k) (_ : k ≤ n), f k) ?_ ?_ ?_ · refine tendsto_nhds_unique ?_ this - refine tendsto_of_tendsto_of_tendsto_of_le_of_le hg₂ tendsto_const_nhds ?_ ?_ - · intro n; rw [← hf₂ n] + refine tendsto_of_tendsto_of_tendsto_of_le_of_le hg₂ tendsto_const_nhds (fun n ↦ ?_) + fun n ↦ ?_ + · rw [← hf₂ n] apply lintegral_mono convert iSup_le_le f n n le_rfl simp only [iSup_apply] - · intro n - exact le_sSup ⟨⨆ (k : ℕ) (_ : k ≤ n), f k, iSup_mem_measurableLE' _ hf₁ _, rfl⟩ + · exact le_sSup ⟨⨆ (k : ℕ) (_ : k ≤ n), f k, iSup_mem_measurableLE' _ hf₁ _, rfl⟩ · intro n refine Measurable.aemeasurable ?_ convert (iSup_mem_measurableLE _ hf₁ n).1 @@ -859,20 +856,24 @@ theorem haveLebesgueDecomposition_of_finiteMeasure [IsFiniteMeasure μ] [IsFinit have hξm : Measurable ξ := by convert Measurable.iSup fun n ↦ (iSup_mem_measurableLE _ hf₁ n).1 simp [hξ] - -- `ξ` is the `f` in the theorem statement and we set `μ₁` to be `μ - ν.withDensity ξ` - -- since we need `μ₁ + ν.withDensity ξ = μ` - set μ₁ := μ - ν.withDensity ξ with hμ₁ + -- we see that `ξ` has the largest integral among all functions in `measurableLE` + have hξle A (hA : MeasurableSet A) : ∫⁻ a in A, ξ a ∂ν ≤ μ A := by + rw [hξ] + simp_rw [iSup_apply] + rw [lintegral_iSup (fun n ↦ (iSup_mem_measurableLE _ hf₁ n).1) (iSup_monotone _)] + exact iSup_le fun n ↦ (iSup_mem_measurableLE _ hf₁ n).2 A hA have hle : ν.withDensity ξ ≤ μ := by - refine le_iff.2 fun B hB ↦ ?_ - rw [hξ, withDensity_apply _ hB] - simp_rw [iSup_apply] - rw [lintegral_iSup (fun i ↦ (iSup_mem_measurableLE _ hf₁ i).1) (iSup_monotone _)] - exact iSup_le fun i ↦ (iSup_mem_measurableLE _ hf₁ i).2 B hB + refine le_intro fun B hB _ ↦ ?_ + rw [withDensity_apply _ hB] + exact hξle B hB have : IsFiniteMeasure (ν.withDensity ξ) := by refine isFiniteMeasure_withDensity ?_ have hle' := hle univ rw [withDensity_apply _ MeasurableSet.univ, Measure.restrict_univ] at hle' exact ne_top_of_le_ne_top (measure_ne_top _ _) hle' + -- `ξ` is the `f` in the theorem statement and we set `μ₁` to be `μ - ν.withDensity ξ` + -- since we need `μ₁ + ν.withDensity ξ = μ` + set μ₁ := μ - ν.withDensity ξ with hμ₁ refine ⟨⟨μ₁, ξ⟩, hξm, ?_, ?_⟩ · by_contra h -- if they are not mutually singular, then from `exists_positive_of_not_mutuallySingular`, @@ -880,35 +881,18 @@ theorem haveLebesgueDecomposition_of_finiteMeasure [IsFiniteMeasure μ] [IsFinit -- positive with respect to `ν - εμ` obtain ⟨ε, hε₁, E, hE₁, hE₂, hE₃⟩ := exists_positive_of_not_mutuallySingular μ₁ ν h simp_rw [hμ₁] at hE₃ - have hξle : ∀ A, MeasurableSet A → (∫⁻ a in A, ξ a ∂ν) ≤ μ A := by - intro A hA; rw [hξ] - simp_rw [iSup_apply] - rw [lintegral_iSup (fun n ↦ (iSup_mem_measurableLE _ hf₁ n).1) (iSup_monotone _)] - exact iSup_le fun n ↦ (iSup_mem_measurableLE _ hf₁ n).2 A hA -- since `E` is positive, we have `∫⁻ a in A ∩ E, ε + ξ a ∂ν ≤ μ (A ∩ E)` for all `A` - have hε₂ : ∀ A : Set α, MeasurableSet A → (∫⁻ a in A ∩ E, ε + ξ a ∂ν) ≤ μ (A ∩ E) := by - intro A hA - have := subset_le_of_restrict_le_restrict _ _ hE₁ hE₃ A.inter_subset_right - rwa [zero_apply, toSignedMeasure_sub_apply (hA.inter hE₁), - Measure.sub_apply (hA.inter hE₁) hle, - ENNReal.toReal_sub_of_le _ (measure_ne_top _ _), sub_nonneg, le_sub_iff_add_le, - ← ENNReal.toReal_add, ENNReal.toReal_le_toReal, Measure.coe_smul, Pi.smul_apply, - withDensity_apply _ (hA.inter hE₁), show ε • ν (A ∩ E) = (ε : ℝ≥0∞) * ν (A ∩ E) by rfl, - ← setLIntegral_const, ← lintegral_add_left measurable_const] at this - · rw [Ne, ENNReal.add_eq_top, not_or] - exact ⟨measure_ne_top _ _, measure_ne_top _ _⟩ - · exact measure_ne_top _ _ - · exact measure_ne_top _ _ - · exact measure_ne_top _ _ - · rw [withDensity_apply _ (hA.inter hE₁)] - exact hξle (A ∩ E) (hA.inter hE₁) + have hε₂ (A : Set α) (hA : MeasurableSet A) : ∫⁻ a in A ∩ E, ε + ξ a ∂ν ≤ μ (A ∩ E) := by + specialize hE₃ A hA + rw [lintegral_add_left measurable_const, lintegral_const, restrict_apply_univ] + rw [Measure.sub_apply (hA.inter hE₁) hle, withDensity_apply _ (hA.inter hE₁)] at hE₃ + refine add_le_of_le_tsub_right_of_le (hξle _ (hA.inter hE₁)) hE₃ -- from this, we can show `ξ + ε * E.indicator` is a function in `measurableLE` with -- integral greater than `ξ` have hξε : (ξ + E.indicator fun _ ↦ (ε : ℝ≥0∞)) ∈ measurableLE ν μ := by - refine ⟨Measurable.add hξm (Measurable.indicator measurable_const hE₁), fun A hA ↦ ?_⟩ - have : - (∫⁻ a in A, (ξ + E.indicator fun _ ↦ (ε : ℝ≥0∞)) a ∂ν) = - (∫⁻ a in A ∩ E, ε + ξ a ∂ν) + ∫⁻ a in A \ E, ξ a ∂ν := by + refine ⟨hξm.add (measurable_const.indicator hE₁), fun A hA ↦ ?_⟩ + have : ∫⁻ a in A, (ξ + E.indicator fun _ ↦ (ε : ℝ≥0∞)) a ∂ν = + ∫⁻ a in A ∩ E, ε + ξ a ∂ν + ∫⁻ a in A \ E, ξ a ∂ν := by simp only [lintegral_add_left measurable_const, lintegral_add_left hξm, setLIntegral_const, add_assoc, lintegral_inter_add_diff _ _ hE₁, Pi.add_apply, lintegral_indicator hE₁, restrict_apply hE₁] @@ -924,9 +908,10 @@ theorem haveLebesgueDecomposition_of_finiteMeasure [IsFiniteMeasure μ] [IsFinit have := measure_ne_top (ν.withDensity ξ) univ rwa [withDensity_apply _ MeasurableSet.univ, Measure.restrict_univ] at this -- since `ν.withDensity ξ ≤ μ`, it is clear that `μ = μ₁ + ν.withDensity ξ` - · rw [hμ₁]; ext1 A hA + · rw [hμ₁] + ext1 A hA rw [Measure.coe_add, Pi.add_apply, Measure.sub_apply hA hle, add_comm, - add_tsub_cancel_of_le (hle A)]⟩ + add_tsub_cancel_of_le (hle A)] /-- If any finite measure has a Lebesgue decomposition with respect to `ν`, then the same is true for any s-finite measure. -/ diff --git a/Mathlib/MeasureTheory/Decomposition/UnsignedHahn.lean b/Mathlib/MeasureTheory/Decomposition/UnsignedHahn.lean index fb69ad9ffd389..c3ebf52d3cb54 100644 --- a/Mathlib/MeasureTheory/Decomposition/UnsignedHahn.lean +++ b/Mathlib/MeasureTheory/Decomposition/UnsignedHahn.lean @@ -26,10 +26,10 @@ open Set Filter Topology ENNReal namespace MeasureTheory -variable {α : Type*} [MeasurableSpace α] {μ ν : Measure α} +variable {α : Type*} {mα : MeasurableSpace α} /-- **Hahn decomposition theorem** -/ -theorem hahn_decomposition [IsFiniteMeasure μ] [IsFiniteMeasure ν] : +theorem hahn_decomposition (μ ν : Measure α) [IsFiniteMeasure μ] [IsFiniteMeasure ν] : ∃ s, MeasurableSet s ∧ (∀ t, MeasurableSet t → t ⊆ s → ν t ≤ μ t) ∧ ∀ t, MeasurableSet t → t ⊆ sᶜ → μ t ≤ ν t := by let d : Set α → ℝ := fun s => ((μ s).toNNReal : ℝ) - (ν s).toNNReal From 0571c82877b7c2b508cbc9a39e041eb117d050b7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Fri, 3 Jan 2025 14:55:08 +0000 Subject: [PATCH 120/189] feat: add fst_compProd_apply (#20429) Extract from the proof of `fst_compProd` a new lemma that applies more generally than Markov kernels. --- .../Probability/Kernel/Composition/Basic.lean | 21 ++++++++++--------- 1 file changed, 11 insertions(+), 10 deletions(-) diff --git a/Mathlib/Probability/Kernel/Composition/Basic.lean b/Mathlib/Probability/Kernel/Composition/Basic.lean index 8fd4daf2faf96..fd06f343aa595 100644 --- a/Mathlib/Probability/Kernel/Composition/Basic.lean +++ b/Mathlib/Probability/Kernel/Composition/Basic.lean @@ -963,19 +963,20 @@ lemma fst_map_id_prod (κ : Kernel α β) {γ : Type*} {mγ : MeasurableSpace γ fst (map κ (fun a ↦ (a, f a))) = κ := by rw [fst_map_prod _ hf, Kernel.map_id'] +/-- If `η` is a Markov kernel, use instead `fst_compProd` to get `(κ ⊗ₖ η).fst = κ`. -/ +lemma fst_compProd_apply (κ : Kernel α β) (η : Kernel (α × β) γ) + [IsSFiniteKernel κ] [IsSFiniteKernel η] (x : α) {s : Set β} (hs : MeasurableSet s) : + (κ ⊗ₖ η).fst x s = ∫⁻ b, s.indicator (fun b ↦ η (x, b) Set.univ) b ∂(κ x) := by + rw [Kernel.fst_apply' _ _ hs, Kernel.compProd_apply] + swap; · exact measurable_fst hs + have h_eq b : η (x, b) {c | b ∈ s} = s.indicator (fun b ↦ η (x, b) Set.univ) b := by + by_cases hb : b ∈ s <;> simp [hb] + simp_rw [Set.mem_setOf_eq, h_eq] + @[simp] lemma fst_compProd (κ : Kernel α β) (η : Kernel (α × β) γ) [IsSFiniteKernel κ] [IsMarkovKernel η] : fst (κ ⊗ₖ η) = κ := by - ext x s hs - rw [fst_apply' _ _ hs, compProd_apply] - swap; · exact measurable_fst hs - simp only [Set.mem_setOf_eq] - classical - have : ∀ b : β, η (x, b) {_c | b ∈ s} = s.indicator (fun _ ↦ 1) b := by - intro b - by_cases hb : b ∈ s <;> simp [hb] - simp_rw [this] - rw [lintegral_indicator_const hs, one_mul] + ext x s hs; simp [fst_compProd_apply, hs] lemma fst_prodMkLeft (δ : Type*) [MeasurableSpace δ] (κ : Kernel α (β × γ)) : fst (prodMkLeft δ κ) = prodMkLeft δ (fst κ) := rfl From fbe1f6b02f2333ea5d47de669bc8e9a365031361 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Fri, 3 Jan 2025 17:09:26 +0000 Subject: [PATCH 121/189] feat: add sumPiEquivProdPi and piUnique (#20195) From FLT. Add topological and algebraic versions of `Equiv.sumPiEquivProdPi` and `Equiv.piUnique`. --- Mathlib/Algebra/Module/Equiv/Basic.lean | 33 ++++++++++++++++++ Mathlib/Topology/Algebra/Module/Equiv.lean | 39 ++++++++++++++++++++++ Mathlib/Topology/Basic.lean | 6 ++++ Mathlib/Topology/Homeomorph.lean | 28 +++++++++++++--- 4 files changed, 101 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/Module/Equiv/Basic.lean b/Mathlib/Algebra/Module/Equiv/Basic.lean index 1ccd71d6cbcbb..a22205a1e0f06 100644 --- a/Mathlib/Algebra/Module/Equiv/Basic.lean +++ b/Mathlib/Algebra/Module/Equiv/Basic.lean @@ -11,6 +11,7 @@ import Mathlib.Algebra.Module.Equiv.Defs import Mathlib.Algebra.Module.Hom import Mathlib.Algebra.Module.LinearMap.End import Mathlib.Algebra.Module.Pi +import Mathlib.Algebra.Module.Prod /-! # Further results on (semi)linear equivalences. @@ -698,4 +699,36 @@ end LinearEquiv end FunLeft +section Pi + +namespace LinearEquiv + +/-- The product over `S ⊕ T` of a family of modules is isomorphic to the product of +(the product over `S`) and (the product over `T`). + +This is `Equiv.sumPiEquivProdPi` as a `LinearEquiv`. +-/ +def sumPiEquivProdPi (R : Type*) [Semiring R] (S T : Type*) (A : S ⊕ T → Type*) + [∀ st, AddCommMonoid (A st)] [∀ st, Module R (A st)] : + (Π (st : S ⊕ T), A st) ≃ₗ[R] (Π (s : S), A (.inl s)) × (Π (t : T), A (.inr t)) where + __ := Equiv.sumPiEquivProdPi _ + map_add' _ _ := rfl + map_smul' _ _ := rfl + +/-- The product `Π t : α, f t` of a family of modules is linearly isomorphic to the module +`f ⬝` when `α` only contains `⬝`. + +This is `Equiv.piUnique` as a `LinearEquiv`. +-/ +@[simps (config := .asFn)] +def piUnique {α : Type*} [Unique α] (R : Type*) [Semiring R] (f : α → Type*) + [∀ x, AddCommMonoid (f x)] [∀ x, Module R (f x)] : (Π t : α, f t) ≃ₗ[R] f default where + __ := Equiv.piUnique _ + map_add' _ _ := rfl + map_smul' _ _ := rfl + +end LinearEquiv + +end Pi + end AddCommMonoid diff --git a/Mathlib/Topology/Algebra/Module/Equiv.lean b/Mathlib/Topology/Algebra/Module/Equiv.lean index b2d1f0f7870e1..fb217bf9ceb00 100644 --- a/Mathlib/Topology/Algebra/Module/Equiv.lean +++ b/Mathlib/Topology/Algebra/Module/Equiv.lean @@ -567,6 +567,45 @@ def arrowCongrEquiv (e₁₂ : M₁ ≃SL[σ₁₂] M₂) (e₄₃ : M₄ ≃SL[ ContinuousLinearMap.ext fun x => by simp only [ContinuousLinearMap.comp_apply, apply_symm_apply, coe_coe] +section Pi + +/-- Combine a family of linear equivalences into a linear equivalence of `pi`-types. +This is `Equiv.piCongrLeft` as a `ContinuousLinearEquiv`. +-/ +def piCongrLeft (R : Type*) [Semiring R] {ι ι' : Type*} + (φ : ι → Type*) [∀ i, AddCommMonoid (φ i)] [∀ i, Module R (φ i)] + [∀ i, TopologicalSpace (φ i)] + (e : ι' ≃ ι) : ((i' : ι') → φ (e i')) ≃L[R] (i : ι) → φ i where + __ := Homeomorph.piCongrLeft e + __ := LinearEquiv.piCongrLeft R φ e + +/-- The product over `S ⊕ T` of a family of topological modules +is isomorphic (topologically and alegbraically) to the product of +(the product over `S`) and (the product over `T`). + +This is `Equiv.sumPiEquivProdPi` as a `ContinuousLinearEquiv`. +-/ +def sumPiEquivProdPi (R : Type*) [Semiring R] (S T : Type*) + (A : S ⊕ T → Type*) [∀ st, AddCommMonoid (A st)] [∀ st, Module R (A st)] + [∀ st, TopologicalSpace (A st)] : + ((st : S ⊕ T) → A st) ≃L[R] ((s : S) → A (Sum.inl s)) × ((t : T) → A (Sum.inr t)) where + __ := LinearEquiv.sumPiEquivProdPi R S T A + __ := Homeomorph.sumPiEquivProdPi S T A + +/-- The product `Π t : α, f t` of a family of topological modules is isomorphic +(both topologically and algebraically) to the space `f ⬝` when `α` only contains `⬝`. + +This is `Equiv.piUnique` as a `ContinuousLinearEquiv`. +-/ +@[simps! (config := .asFn)] +def piUnique {α : Type*} [Unique α] (R : Type*) [Semiring R] (f : α → Type*) + [∀ x, AddCommMonoid (f x)] [∀ x, Module R (f x)] [∀ x, TopologicalSpace (f x)] : + (Π t, f t) ≃L[R] f default where + __ := LinearEquiv.piUnique R f + __ := Homeomorph.piUnique f + +end Pi + section piCongrRight variable {ι : Type*} {M : ι → Type*} [∀ i, TopologicalSpace (M i)] [∀ i, AddCommMonoid (M i)] diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index 76073103460ca..e6f284b71d21c 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -1365,6 +1365,12 @@ theorem IsOpen.preimage (hf : Continuous f) {t : Set Y} (h : IsOpen t) : IsOpen (f ⁻¹' t) := hf.isOpen_preimage t h +lemma Equiv.continuous_symm_iff (e : X ≃ Y) : Continuous e.symm ↔ IsOpenMap e := by + simp_rw [continuous_def, ← Set.image_equiv_eq_preimage_symm, IsOpenMap] + +lemma Equiv.isOpenMap_symm_iff (e : X ≃ Y) : IsOpenMap e.symm ↔ Continuous e := by + simp_rw [← Equiv.continuous_symm_iff, Equiv.symm_symm] + theorem continuous_congr {g : X → Y} (h : ∀ x, f x = g x) : Continuous f ↔ Continuous g := .of_eq <| congrArg _ <| funext h diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index f6859612b043c..224fd0408685c 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -427,11 +427,7 @@ theorem locallyCompactSpace_iff (h : X ≃ₜ Y) : @[simps toEquiv] def homeomorphOfContinuousOpen (e : X ≃ Y) (h₁ : Continuous e) (h₂ : IsOpenMap e) : X ≃ₜ Y where continuous_toFun := h₁ - continuous_invFun := by - rw [continuous_def] - intro s hs - convert ← h₂ s hs using 1 - apply e.image_eq_preimage + continuous_invFun := e.continuous_symm_iff.2 h₂ toEquiv := e /-- If a bijective map `e : X ≃ Y` is continuous and closed, then it is a homeomorphism. -/ @@ -671,6 +667,28 @@ def homeomorphOfUnique [Unique X] [Unique Y] : X ≃ₜ Y := continuous_toFun := continuous_const continuous_invFun := continuous_const } +/-- The product over `S ⊕ T` of a family of topological spaces +is homeomorphic to the product of (the product over `S`) and (the product over `T`). + +This is `Equiv.sumPiEquivProdPi` as a `Homeomorph`. +-/ +def sumPiEquivProdPi (S T : Type*) (A : S ⊕ T → Type*) + [∀ st, TopologicalSpace (A st)] : + (Π (st : S ⊕ T), A st) ≃ₜ (Π (s : S), A (.inl s)) × (Π (t : T), A (.inr t)) where + __ := Equiv.sumPiEquivProdPi _ + continuous_toFun := Continuous.prod_mk (by fun_prop) (by fun_prop) + continuous_invFun := continuous_pi <| by rintro (s | t) <;> simp <;> fun_prop + +/-- The product `Π t : α, f t` of a family of topological spaces is homeomorphic to the +space `f ⬝` when `α` only contains `⬝`. + +This is `Equiv.piUnique` as a `Homeomorph`. +-/ +@[simps! (config := .asFn)] +def piUnique {α : Type*} [Unique α] (f : α → Type*) [∀ x, TopologicalSpace (f x)] : + (Π t, f t) ≃ₜ f default := + homeomorphOfContinuousOpen (Equiv.piUnique f) (continuous_apply default) (isOpenMap_eval _) + end prod /-- `Equiv.piCongrLeft` as a homeomorphism: this is the natural homeomorphism From 2e72e436be9e6c0cd6c17654aba07b283b87344b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 3 Jan 2025 18:26:13 +0000 Subject: [PATCH 122/189] chore(1000): remove `identifiers` (#20445) This key is not recognized by CI --- docs/1000.yaml | 1 - 1 file changed, 1 deletion(-) diff --git a/docs/1000.yaml b/docs/1000.yaml index 27d543c4435d4..64168a6a6f4ae 100644 --- a/docs/1000.yaml +++ b/docs/1000.yaml @@ -1066,7 +1066,6 @@ Q1050203: Q1050932: title: Hartogs's theorem url: https://github.com/girving/ray/blob/main/Ray/Hartogs/Hartogs.lean - identifiers: Pair.hartogs author: Geoffrey Irving Q1051404: From e4220ebf1cdae020e7e79ba5a35d26e163bac828 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 3 Jan 2025 18:41:13 +0000 Subject: [PATCH 123/189] doc(Algebra/Lie/Weights/Basic): fix typo (#20450) --- Mathlib/Algebra/Lie/Weights/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index d6c6f9d9c417c..7ecb88285cfc7 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -709,7 +709,7 @@ noncomputable instance Weight.instFintype [NoZeroSMulDivisors R M] [IsNoetherian Fintype (Weight R L M) := Fintype.ofFinite _ -/-- A Lie module `M` of a Lie algebra `L` is triangularizable if the endomorhpism of `M` defined by +/-- A Lie module `M` of a Lie algebra `L` is triangularizable if the endomorphism of `M` defined by any `x : L` is triangularizable. -/ class IsTriangularizable : Prop where maxGenEigenspace_eq_top : ∀ x, ⨆ φ, (toEnd R L M x).maxGenEigenspace φ = ⊤ From 155a0a990b89951fc9c745bbd32f1430cfc2a149 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Fri, 3 Jan 2025 22:22:58 +0000 Subject: [PATCH 124/189] feat: add `ENNReal.finStronglyMeasurable_of_measurable` (#20404) A measurable function with finite Lebesgue integral can be approximated by simple functions whose support has finite measure. I created a new file StronglyMeasurable/ENNReal instead of using StronglyMeasurable/Lemmas because those lemmas will be useful in places that should not import things like the Bochner integral. --- Mathlib.lean | 1 + .../MeasureTheory/Function/SimpleFunc.lean | 19 +++++++++++ .../Function/SimpleFuncDenseLp.lean | 7 ---- .../Function/StronglyMeasurable/ENNReal.lean | 34 +++++++++++++++++++ Mathlib/MeasureTheory/Integral/Lebesgue.lean | 11 ++++++ 5 files changed, 65 insertions(+), 7 deletions(-) create mode 100644 Mathlib/MeasureTheory/Function/StronglyMeasurable/ENNReal.lean diff --git a/Mathlib.lean b/Mathlib.lean index baddcd1b0a810..04ae5f192f56b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3698,6 +3698,7 @@ import Mathlib.MeasureTheory.Function.SpecialFunctions.Basic import Mathlib.MeasureTheory.Function.SpecialFunctions.Inner import Mathlib.MeasureTheory.Function.SpecialFunctions.RCLike import Mathlib.MeasureTheory.Function.StronglyMeasurable.Basic +import Mathlib.MeasureTheory.Function.StronglyMeasurable.ENNReal import Mathlib.MeasureTheory.Function.StronglyMeasurable.Inner import Mathlib.MeasureTheory.Function.StronglyMeasurable.Lemmas import Mathlib.MeasureTheory.Function.StronglyMeasurable.Lp diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index f8ef14327a38d..05df7e2bf5776 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -779,6 +779,12 @@ theorem eapprox_comp [MeasurableSpace γ] {f : γ → ℝ≥0∞} {g : α → γ (hg : Measurable g) : (eapprox (f ∘ g) n : α → ℝ≥0∞) = (eapprox f n : γ →ₛ ℝ≥0∞) ∘ g := funext fun a => approx_comp a hf hg +lemma tendsto_eapprox {f : α → ℝ≥0∞} (hf_meas : Measurable f) (a : α) : + Tendsto (fun n ↦ eapprox f n a) atTop (𝓝 (f a)) := by + nth_rw 2 [← iSup_coe_eapprox hf_meas] + rw [iSup_apply] + exact tendsto_atTop_iSup fun _ _ hnm ↦ monotone_eapprox f hnm a + /-- Approximate a function `α → ℝ≥0∞` by a series of simple functions taking their values in `ℝ≥0`. -/ def eapproxDiff (f : α → ℝ≥0∞) : ℕ → α →ₛ ℝ≥0 @@ -1013,6 +1019,13 @@ theorem measurableSet_support [MeasurableSpace α] (f : α →ₛ β) : Measurab rw [f.support_eq] exact Finset.measurableSet_biUnion _ fun y _ => measurableSet_fiber _ _ +lemma measure_support_lt_top (f : α →ₛ β) (hf : ∀ y, y ≠ 0 → μ (f ⁻¹' {y}) < ∞) : + μ (support f) < ∞ := by + rw [support_eq] + refine (measure_biUnion_finset_le _ _).trans_lt (ENNReal.sum_lt_top.mpr fun y hy => ?_) + rw [Finset.mem_filter] at hy + exact hf y hy.2 + /-- A `SimpleFunc` has finite measure support if it is equal to `0` outside of a set of finite measure. -/ protected def FinMeasSupp {_m : MeasurableSpace α} (f : α →ₛ β) (μ : Measure α) : Prop := @@ -1090,6 +1103,12 @@ theorem iff_lintegral_lt_top {f : α →ₛ ℝ≥0∞} (hf : ∀ᵐ a ∂μ, f end FinMeasSupp +lemma measure_support_lt_top_of_lintegral_ne_top {f : α →ₛ ℝ≥0∞} (hf : f.lintegral μ ≠ ∞) : + μ (support f) < ∞ := by + refine measure_support_lt_top f ?_ + rw [← finMeasSupp_iff] + exact FinMeasSupp.of_lintegral_ne_top hf + end FinMeasSupp /-- To prove something for an arbitrary simple function, it suffices to show diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index b8a67af0bd1c9..b8cf7ba03d100 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -366,13 +366,6 @@ theorem measure_preimage_lt_top_of_integrable (f : α →ₛ E) (hf : Integrable (hx : x ≠ 0) : μ (f ⁻¹' {x}) < ∞ := integrable_iff.mp hf x hx -theorem measure_support_lt_top [Zero β] (f : α →ₛ β) (hf : ∀ y, y ≠ 0 → μ (f ⁻¹' {y}) < ∞) : - μ (support f) < ∞ := by - rw [support_eq] - refine (measure_biUnion_finset_le _ _).trans_lt (ENNReal.sum_lt_top.mpr fun y hy => ?_) - rw [Finset.mem_filter] at hy - exact hf y hy.2 - theorem measure_support_lt_top_of_memℒp (f : α →ₛ E) (hf : Memℒp f p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : μ (support f) < ∞ := f.measure_support_lt_top ((memℒp_iff hp_ne_zero hp_ne_top).mp hf) diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/ENNReal.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/ENNReal.lean new file mode 100644 index 0000000000000..98bd066bdea2c --- /dev/null +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/ENNReal.lean @@ -0,0 +1,34 @@ +/- +Copyright (c) 2025 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.Function.StronglyMeasurable.Basic +import Mathlib.MeasureTheory.Integral.Lebesgue + +/-! +# Finitely strongly measurable functions with value in ENNReal + +A measurable function with finite Lebesgue integral can be approximated by simple functions +whose support has finite measure. + +-/ + +open MeasureTheory +open scoped ENNReal + +variable {α : Type*} {m : MeasurableSpace α} {μ : Measure α} {f : α → ℝ≥0∞} + +lemma ENNReal.finStronglyMeasurable_of_measurable (hf : ∫⁻ x, f x ∂μ ≠ ∞) + (hf_meas : Measurable f) : + FinStronglyMeasurable f μ := + ⟨SimpleFunc.eapprox f, measure_support_eapprox_lt_top hf_meas hf, + SimpleFunc.tendsto_eapprox hf_meas⟩ + +lemma ENNReal.aefinStronglyMeasurable_of_aemeasurable (hf : ∫⁻ x, f x ∂μ ≠ ∞) + (hf_meas : AEMeasurable f μ) : + AEFinStronglyMeasurable f μ := by + refine ⟨hf_meas.mk f, ENNReal.finStronglyMeasurable_of_measurable ?_ hf_meas.measurable_mk, + hf_meas.ae_eq_mk⟩ + rwa [lintegral_congr_ae hf_meas.ae_eq_mk.symm] diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index e5d534014c897..71735659713fe 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -449,6 +449,17 @@ theorem lintegral_eq_iSup_eapprox_lintegral {f : α → ℝ≥0∞} (hf : Measur _ = ⨆ n, (eapprox f n).lintegral μ := by congr; ext n; rw [(eapprox f n).lintegral_eq_lintegral] +lemma lintegral_eapprox_le_lintegral {f : α → ℝ≥0∞} (hf : Measurable f) (n : ℕ) : + (eapprox f n).lintegral μ ≤ ∫⁻ x, f x ∂μ := by + rw [lintegral_eq_iSup_eapprox_lintegral hf] + exact le_iSup (fun n ↦ (eapprox f n).lintegral μ) n + +lemma measure_support_eapprox_lt_top {f : α → ℝ≥0∞} (hf_meas : Measurable f) + (hf : ∫⁻ x, f x ∂μ ≠ ∞) (n : ℕ) : + μ (support (eapprox f n)) < ∞ := + measure_support_lt_top_of_lintegral_ne_top <| + ((lintegral_eapprox_le_lintegral hf_meas n).trans_lt hf.lt_top).ne + /-- If `f` has finite integral, then `∫⁻ x in s, f x ∂μ` is absolutely continuous in `s`: it tends to zero as `μ s` tends to zero. This lemma states this fact in terms of `ε` and `δ`. -/ theorem exists_pos_setLIntegral_lt_of_measure_lt {f : α → ℝ≥0∞} (h : ∫⁻ x, f x ∂μ ≠ ∞) {ε : ℝ≥0∞} From 166dc0de051a344cb02d7c2351b41717f0e800b7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 3 Jan 2025 23:07:04 +0000 Subject: [PATCH 125/189] =?UTF-8?q?feat(Polynomial):=20`(C=20a=20*=20p).de?= =?UTF-8?q?gree=20=3D=20p.degree`=20if=20`a=20*=20p.leadingCoeff=20?= =?UTF-8?q?=E2=89=A0=200`=20(#20446)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From GrowthInGroups (LeanCamCombi) Co-authored-by: Andrew Yang --- Mathlib/Algebra/Polynomial/Degree/Lemmas.lean | 26 ++++++++++--------- Mathlib/Algebra/Polynomial/Taylor.lean | 2 +- 2 files changed, 15 insertions(+), 13 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean b/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean index 6f001cb5e20e7..c6557963d2e95 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean @@ -88,17 +88,13 @@ theorem natDegree_add_le_iff_right {n : ℕ} (p q : R[X]) (pn : p.natDegree ≤ rw [add_comm] exact natDegree_add_le_iff_left _ _ pn -theorem natDegree_C_mul_le (a : R) (f : R[X]) : (C a * f).natDegree ≤ f.natDegree := - calc - (C a * f).natDegree ≤ (C a).natDegree + f.natDegree := natDegree_mul_le - _ = 0 + f.natDegree := by rw [natDegree_C a] - _ = f.natDegree := zero_add _ - -theorem natDegree_mul_C_le (f : R[X]) (a : R) : (f * C a).natDegree ≤ f.natDegree := - calc - (f * C a).natDegree ≤ f.natDegree + (C a).natDegree := natDegree_mul_le - _ = f.natDegree + 0 := by rw [natDegree_C a] - _ = f.natDegree := add_zero _ +-- TODO: Do we really want the following two lemmas? They are straightforward consequences of a +-- more atomic lemma +theorem natDegree_C_mul_le (a : R) (f : R[X]) : (C a * f).natDegree ≤ f.natDegree := by + simpa using natDegree_mul_le (p := C a) + +theorem natDegree_mul_C_le (f : R[X]) (a : R) : (f * C a).natDegree ≤ f.natDegree := by + simpa using natDegree_mul_le (q := C a) theorem eq_natDegree_of_le_mem_support (pn : p.natDegree ≤ n) (ns : n ∈ p.support) : p.natDegree = n := @@ -132,12 +128,18 @@ theorem natDegree_mul_C_eq_of_mul_ne_zero (h : p.leadingCoeff * a ≠ 0) : /-- Although not explicitly stated, the assumptions of lemma `nat_degree_C_mul_eq_of_mul_ne_zero` force the polynomial `p` to be non-zero, via `p.leading_coeff ≠ 0`. -/ -theorem natDegree_C_mul_eq_of_mul_ne_zero (h : a * p.leadingCoeff ≠ 0) : +theorem natDegree_C_mul_of_mul_ne_zero (h : a * p.leadingCoeff ≠ 0) : (C a * p).natDegree = p.natDegree := by refine eq_natDegree_of_le_mem_support (natDegree_C_mul_le a p) ?_ refine mem_support_iff.mpr ?_ rwa [coeff_C_mul] +@[deprecated (since := "2025-01-03")] +alias natDegree_C_mul_eq_of_mul_ne_zero := natDegree_C_mul_of_mul_ne_zero + +lemma degree_C_mul_of_mul_ne_zero (h : a * p.leadingCoeff ≠ 0) : (C a * p).degree = p.degree := by + rw [degree_mul' (by simpa)]; simp [left_ne_zero_of_mul h] + theorem natDegree_add_coeff_mul (f g : R[X]) : (f * g).coeff (f.natDegree + g.natDegree) = f.coeff f.natDegree * g.coeff g.natDegree := by simp only [coeff_natDegree, coeff_mul_degree_add_degree] diff --git a/Mathlib/Algebra/Polynomial/Taylor.lean b/Mathlib/Algebra/Polynomial/Taylor.lean index f3f2c2b9b9fd7..071da34ab0b61 100644 --- a/Mathlib/Algebra/Polynomial/Taylor.lean +++ b/Mathlib/Algebra/Polynomial/Taylor.lean @@ -83,7 +83,7 @@ theorem natDegree_taylor (p : R[X]) (r : R) : natDegree (taylor r p) = natDegree refine map_natDegree_eq_natDegree _ ?_ nontriviality R intro n c c0 - simp [taylor_monomial, natDegree_C_mul_eq_of_mul_ne_zero, natDegree_pow_X_add_C, c0] + simp [taylor_monomial, natDegree_C_mul_of_mul_ne_zero, natDegree_pow_X_add_C, c0] @[simp] theorem taylor_mul {R} [CommSemiring R] (r : R) (p q : R[X]) : From f87a13504e7b6346768d33065064188a8af4c084 Mon Sep 17 00:00:00 2001 From: Artie Khovanov Date: Fri, 3 Jan 2025 23:29:25 +0000 Subject: [PATCH 126/189] chore(Algebra): Improve attribute generation (#20451) * Teach `to_additive` about `oneLE` -> `nonneg` and `square` -> `even` * Make `simps` lemma generated for the `carrier` projection consistently `coe_*` * Clarify some documentation Moves: * `\*_coe` -> `coe_\*` where generated by `@[simps]` * `IsOpenMap.functor_obj_coe` -> `IsOpenMap.coe_functor_obj` * `AlgebraicGeometry.Scheme.Hom.opensRange_coe` -> `AlgebraicGeometry.Scheme.Hom.coe_opensRange` --- Mathlib/Algebra/Group/Subsemigroup/Defs.lean | 4 ++-- Mathlib/Algebra/Order/Group/Cone.lean | 15 ++++++++------- Mathlib/Algebra/Order/Monoid/Submonoid.lean | 6 +++--- Mathlib/AlgebraicGeometry/Cover/Open.lean | 2 +- .../AlgebraicGeometry/Morphisms/Separated.lean | 4 ++-- Mathlib/AlgebraicGeometry/Noetherian.lean | 2 +- Mathlib/AlgebraicGeometry/OpenImmersion.lean | 2 +- Mathlib/AlgebraicGeometry/RationalMap.lean | 2 +- Mathlib/AlgebraicGeometry/Restrict.lean | 4 ++-- Mathlib/Geometry/RingedSpace/OpenImmersion.lean | 2 +- .../RingedSpace/PresheafedSpace/Gluing.lean | 4 ++-- Mathlib/ModelTheory/Substructures.lean | 2 +- Mathlib/Order/UpperLower/Basic.lean | 4 ++-- Mathlib/Tactic/ToAdditive/Frontend.lean | 8 +++++--- Mathlib/Topology/Category/TopCat/Opens.lean | 2 +- Mathlib/Topology/Sets/Closeds.lean | 2 +- Mathlib/Topology/Sets/Compacts.lean | 8 ++++---- Mathlib/Topology/Sets/Opens.lean | 2 +- Mathlib/Topology/Sets/Order.lean | 2 +- 19 files changed, 40 insertions(+), 37 deletions(-) diff --git a/Mathlib/Algebra/Group/Subsemigroup/Defs.lean b/Mathlib/Algebra/Group/Subsemigroup/Defs.lean index 53d6f62610fbc..ef136ff0a76ab 100644 --- a/Mathlib/Algebra/Group/Subsemigroup/Defs.lean +++ b/Mathlib/Algebra/Group/Subsemigroup/Defs.lean @@ -97,8 +97,8 @@ instance : SetLike (Subsemigroup M) M := @[to_additive] instance : MulMemClass (Subsemigroup M) M where mul_mem := fun {_ _ _} => Subsemigroup.mul_mem' _ -initialize_simps_projections Subsemigroup (carrier → coe) -initialize_simps_projections AddSubsemigroup (carrier → coe) +initialize_simps_projections Subsemigroup (carrier → coe, as_prefix coe) +initialize_simps_projections AddSubsemigroup (carrier → coe, as_prefix coe) @[to_additive (attr := simp)] theorem mem_carrier {s : Subsemigroup M} {x : M} : x ∈ s.carrier ↔ x ∈ s := diff --git a/Mathlib/Algebra/Order/Group/Cone.lean b/Mathlib/Algebra/Order/Group/Cone.lean index 5f110e86b0c0e..d9236d10622d2 100644 --- a/Mathlib/Algebra/Order/Group/Cone.lean +++ b/Mathlib/Algebra/Order/Group/Cone.lean @@ -58,6 +58,9 @@ instance GroupCone.instGroupConeClass (G : Type*) [CommGroup G] : one_mem {C} := C.one_mem' eq_one_of_mem_of_inv_mem {C} := C.eq_one_of_mem_of_inv_mem' +initialize_simps_projections GroupCone (carrier → coe, as_prefix coe) +initialize_simps_projections AddGroupCone (carrier → coe, as_prefix coe) + /-- Typeclass for maximal additive cones. -/ class IsMaxCone {S G : Type*} [AddCommGroup G] [SetLike S G] (C : S) : Prop where mem_or_neg_mem (a : G) : a ∈ C ∨ -a ∈ C @@ -74,19 +77,17 @@ namespace GroupCone variable {H : Type*} [OrderedCommGroup H] {a : H} variable (H) in -/-- Construct a cone from the set of elements of -a partially ordered abelian group that are at least 1. -/ -@[to_additive nonneg -"Construct a cone from the set of non-negative elements of a partially ordered abelian group."] +/-- The cone of elements that are at least 1. -/ +@[to_additive "The cone of non-negative elements."] def oneLE : GroupCone H where __ := Submonoid.oneLE H eq_one_of_mem_of_inv_mem' {a} := by simpa using ge_antisymm -@[to_additive (attr := simp) nonneg_toAddSubmonoid] +@[to_additive (attr := simp)] lemma oneLE_toSubmonoid : (oneLE H).toSubmonoid = .oneLE H := rfl -@[to_additive (attr := simp) mem_nonneg] +@[to_additive (attr := simp)] lemma mem_oneLE : a ∈ oneLE H ↔ 1 ≤ a := Iff.rfl -@[to_additive (attr := simp, norm_cast) coe_nonneg] +@[to_additive (attr := simp, norm_cast)] lemma coe_oneLE : oneLE H = {x : H | 1 ≤ x} := rfl @[to_additive nonneg.isMaxCone] diff --git a/Mathlib/Algebra/Order/Monoid/Submonoid.lean b/Mathlib/Algebra/Order/Monoid/Submonoid.lean index 7961b103fbe1c..f12cf9b2fc99b 100644 --- a/Mathlib/Algebra/Order/Monoid/Submonoid.lean +++ b/Mathlib/Algebra/Order/Monoid/Submonoid.lean @@ -91,8 +91,8 @@ section Preorder variable (M) variable [Monoid M] [Preorder M] [MulLeftMono M] {a : M} -/-- The submonoid of elements greater than `1`. -/ -@[to_additive (attr := simps) nonneg "The submonoid of nonnegative elements."] +/-- The submonoid of elements that are at least `1`. -/ +@[to_additive (attr := simps) "The submonoid of nonnegative elements."] def oneLE : Submonoid M where carrier := Set.Ici 1 mul_mem' := one_le_mul @@ -100,7 +100,7 @@ def oneLE : Submonoid M where variable {M} -@[to_additive (attr := simp) mem_nonneg] lemma mem_oneLE : a ∈ oneLE M ↔ 1 ≤ a := Iff.rfl +@[to_additive (attr := simp)] lemma mem_oneLE : a ∈ oneLE M ↔ 1 ≤ a := Iff.rfl end Preorder end Submonoid diff --git a/Mathlib/AlgebraicGeometry/Cover/Open.lean b/Mathlib/AlgebraicGeometry/Cover/Open.lean index c166e4cf0b736..7c743cdacbf51 100644 --- a/Mathlib/AlgebraicGeometry/Cover/Open.lean +++ b/Mathlib/AlgebraicGeometry/Cover/Open.lean @@ -210,7 +210,7 @@ lemma OpenCover.ext_elem {X : Scheme.{u}} {U : X.Opens} (f g : Γ(X, U)) (𝒰 : fapply TopCat.Sheaf.eq_of_locally_eq' X.sheaf (fun i ↦ (𝒰.map (𝒰.f i)).opensRange ⊓ U) _ (fun _ ↦ homOfLE inf_le_right) · intro x hx - simp only [Opens.iSup_mk, Opens.carrier_eq_coe, Opens.coe_inf, Hom.opensRange_coe, Opens.coe_mk, + simp only [Opens.iSup_mk, Opens.carrier_eq_coe, Opens.coe_inf, Hom.coe_opensRange, Opens.coe_mk, Set.mem_iUnion, Set.mem_inter_iff, Set.mem_range, SetLike.mem_coe, exists_and_right] refine ⟨?_, hx⟩ simpa using ⟨_, 𝒰.covers x⟩ diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean index 516521c7a6539..6c6b1821adf90 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean @@ -126,7 +126,7 @@ lemma Scheme.Pullback.diagonalCoverDiagonalRange_eq_top_of_injective rw [← top_le_iff] rintro x - simp only [diagonalCoverDiagonalRange, openCoverOfBase_J, openCoverOfBase_obj, - openCoverOfLeftRight_J, Opens.iSup_mk, Opens.carrier_eq_coe, Hom.opensRange_coe, Opens.coe_mk, + openCoverOfLeftRight_J, Opens.iSup_mk, Opens.carrier_eq_coe, Hom.coe_opensRange, Opens.coe_mk, Set.mem_iUnion, Set.mem_range, Sigma.exists] have H : (pullback.fst f f).base x = (pullback.snd f f).base x := hf (by rw [← Scheme.comp_base_apply, ← Scheme.comp_base_apply, pullback.condition]) @@ -147,7 +147,7 @@ lemma Scheme.Pullback.range_diagonal_subset_diagonalCoverDiagonalRange : Set.range (pullback.diagonal f).base ⊆ diagonalCoverDiagonalRange f 𝒰 𝒱 := by rintro _ ⟨x, rfl⟩ simp only [diagonalCoverDiagonalRange, openCoverOfBase_J, openCoverOfBase_obj, - openCoverOfLeftRight_J, Opens.iSup_mk, Opens.carrier_eq_coe, Hom.opensRange_coe, Opens.coe_mk, + openCoverOfLeftRight_J, Opens.iSup_mk, Opens.carrier_eq_coe, Hom.coe_opensRange, Opens.coe_mk, Set.mem_iUnion, Set.mem_range, Sigma.exists] let i := 𝒰.f (f.base x) obtain ⟨y : 𝒰.obj i, hy : (𝒰.map i).base y = f.base x⟩ := 𝒰.covers (f.base x) diff --git a/Mathlib/AlgebraicGeometry/Noetherian.lean b/Mathlib/AlgebraicGeometry/Noetherian.lean index a88d237043337..6fbbeea98d0dc 100644 --- a/Mathlib/AlgebraicGeometry/Noetherian.lean +++ b/Mathlib/AlgebraicGeometry/Noetherian.lean @@ -158,7 +158,7 @@ lemma isLocallyNoetherian_of_isOpenImmersion {Y : Scheme} (f : X ⟶ Y) [IsOpenI · suffices Scheme.Hom.opensRange f ⊓ V = V by rw [this] rw [← Opens.coe_inj] - rw [Opens.coe_inf, Scheme.Hom.opensRange_coe, IsOpenMap.functor_obj_coe, + rw [Opens.coe_inf, Scheme.Hom.coe_opensRange, IsOpenMap.coe_functor_obj, Set.inter_eq_right, Set.image_subset_iff, Set.preimage_range] exact Set.subset_univ _ diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index 377bcde76e43c..0c8626614a15c 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -543,7 +543,7 @@ theorem range_pullback_to_base_of_left : Set.range f.base ∩ Set.range g.base := by rw [pullback.condition, Scheme.comp_base, TopCat.coe_comp, Set.range_comp, range_pullback_snd_of_left, Opens.carrier_eq_coe, Opens.map_obj, Opens.coe_mk, - Set.image_preimage_eq_inter_range, Opens.carrier_eq_coe, Scheme.Hom.opensRange_coe] + Set.image_preimage_eq_inter_range, Opens.carrier_eq_coe, Scheme.Hom.coe_opensRange] theorem range_pullback_to_base_of_right : Set.range (pullback.fst g f ≫ g).base = diff --git a/Mathlib/AlgebraicGeometry/RationalMap.lean b/Mathlib/AlgebraicGeometry/RationalMap.lean index 7e848dcf6666c..82168500fbb4a 100644 --- a/Mathlib/AlgebraicGeometry/RationalMap.lean +++ b/Mathlib/AlgebraicGeometry/RationalMap.lean @@ -247,7 +247,7 @@ lemma equiv_of_fromSpecStalkOfMem_eq [IrreducibleSpace X] ((Set.image_subset_range _ _).trans_eq (Subtype.range_val)).trans inf_le_left, ((Set.image_subset_range _ _).trans_eq (Subtype.range_val)).trans inf_le_right, ?_⟩ rw [← cancel_epi (Scheme.Hom.isoImage _ _).hom] - simp only [TopologicalSpace.Opens.carrier_eq_coe, IsOpenMap.functor_obj_coe, + simp only [TopologicalSpace.Opens.carrier_eq_coe, IsOpenMap.coe_functor_obj, TopologicalSpace.Opens.coe_inf, restrict_hom, ← Category.assoc] at e ⊢ convert e using 2 <;> rw [← cancel_mono (Scheme.Opens.ι _)] <;> simp · rw [← f.fromSpecStalkOfMem_restrict hdense inf_le_left ⟨hxf, hxg⟩, diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index b0cb9b5db318c..afcdc6496a7fb 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -233,7 +233,7 @@ theorem Scheme.homOfLE_apply {U V : X.Opens} (e : U ≤ V) (x : U) : theorem Scheme.ι_image_homOfLE_le_ι_image {U V : X.Opens} (e : U ≤ V) (W : Opens V) : U.ι ''ᵁ (X.homOfLE e ⁻¹ᵁ W) ≤ V.ι ''ᵁ W := by - simp only [← SetLike.coe_subset_coe, IsOpenMap.functor_obj_coe, Set.image_subset_iff, + simp only [← SetLike.coe_subset_coe, IsOpenMap.coe_functor_obj, Set.image_subset_iff, Scheme.homOfLE_base, Opens.map_coe, Opens.inclusion'_apply] rintro _ h exact ⟨_, h, rfl⟩ @@ -595,7 +595,7 @@ def morphismRestrictRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : refine Arrow.isoMk' _ _ ((Scheme.Opens.ι _).isoImage _ ≪≫ Scheme.isoOfEq _ ?_) ((Scheme.Opens.ι _).isoImage _) ?_ · ext x - simp only [IsOpenMap.functor_obj_coe, Opens.coe_inclusion', + simp only [IsOpenMap.coe_functor_obj, Opens.coe_inclusion', Opens.map_coe, Set.mem_image, Set.mem_preimage, SetLike.mem_coe, morphismRestrict_base] constructor · rintro ⟨⟨a, h₁⟩, h₂, rfl⟩ diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index 697a19f7df281..9456c6bb4f9f8 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -156,7 +156,7 @@ 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, TopCat.coe_comp] + dsimp only [Opens.map_coe, IsOpenMap.coe_functor_obj, comp_base, TopCat.coe_comp] rw [Set.image_comp, Set.preimage_image_eq _ hg.base_open.injective] /-- For an open immersion `f : X ⟶ Y` and an open set `U ⊆ X`, we have the map `X(U) ⟶ Y(U)`. -/ diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index df614fe6a7795..77f139367118c 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -242,7 +242,7 @@ theorem ι_image_preimage_eq (i j : D.J) (U : Opens (D.U i).carrier) : (opensFunctor (D.f j i)).obj ((Opens.map (𝖣.t j i).base).obj ((Opens.map (𝖣.f i j).base).obj U)) := by ext1 - dsimp only [Opens.map_coe, IsOpenMap.functor_obj_coe] + dsimp only [Opens.map_coe, IsOpenMap.coe_functor_obj] rw [← show _ = (𝖣.ι i).base from 𝖣.ι_gluedIso_inv (PresheafedSpace.forget _) i, ← show _ = (𝖣.ι j).base from 𝖣.ι_gluedIso_inv (PresheafedSpace.forget _) j] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11224): change `rw` to `erw` on `coe_comp` @@ -328,7 +328,7 @@ def ιInvAppπApp {i : D.J} (U : Opens (D.U i).carrier) (j) : D.opensImagePreimageMap i j U ≫ (D.f j k).c.app _ ≫ (D.V (j, k)).presheaf.map (eqToHom ?_) rw [Functor.op_obj] congr 1; ext1 - dsimp only [Functor.op_obj, Opens.map_coe, unop_op, IsOpenMap.functor_obj_coe] + dsimp only [Functor.op_obj, Opens.map_coe, unop_op, IsOpenMap.coe_functor_obj] rw [Set.preimage_preimage] change (D.f j k ≫ 𝖣.ι j).base ⁻¹' _ = _ -- Porting note: used to be `congr 3` diff --git a/Mathlib/ModelTheory/Substructures.lean b/Mathlib/ModelTheory/Substructures.lean index a3b0c3233492f..8b6541e29a94e 100644 --- a/Mathlib/ModelTheory/Substructures.lean +++ b/Mathlib/ModelTheory/Substructures.lean @@ -106,7 +106,7 @@ instance instSetLike : SetLike (L.Substructure M) M := def Simps.coe (S : L.Substructure M) : Set M := S -initialize_simps_projections Substructure (carrier → coe) +initialize_simps_projections Substructure (carrier → coe, as_prefix coe) @[simp] theorem mem_carrier {s : L.Substructure M} {x : M} : x ∈ s.carrier ↔ x ∈ s := diff --git a/Mathlib/Order/UpperLower/Basic.lean b/Mathlib/Order/UpperLower/Basic.lean index dbe86c2cd1521..83d0d3803a83b 100644 --- a/Mathlib/Order/UpperLower/Basic.lean +++ b/Mathlib/Order/UpperLower/Basic.lean @@ -429,7 +429,7 @@ instance : SetLike (UpperSet α) α where /-- See Note [custom simps projection]. -/ def Simps.coe (s : UpperSet α) : Set α := s -initialize_simps_projections UpperSet (carrier → coe) +initialize_simps_projections UpperSet (carrier → coe, as_prefix coe) @[ext] theorem ext {s t : UpperSet α} : (s : Set α) = t → s = t := @@ -455,7 +455,7 @@ instance : SetLike (LowerSet α) α where /-- See Note [custom simps projection]. -/ def Simps.coe (s : LowerSet α) : Set α := s -initialize_simps_projections LowerSet (carrier → coe) +initialize_simps_projections LowerSet (carrier → coe, as_prefix coe) @[ext] theorem ext {s t : LowerSet α} : (s : Set α) = t → s = t := diff --git a/Mathlib/Tactic/ToAdditive/Frontend.lean b/Mathlib/Tactic/ToAdditive/Frontend.lean index 108924443e60e..38471ca5c0c9b 100644 --- a/Mathlib/Tactic/ToAdditive/Frontend.lean +++ b/Mathlib/Tactic/ToAdditive/Frontend.lean @@ -977,6 +977,7 @@ def nameDict : String → List String | "multipliable"=> ["summable"] | "gpfree" => ["apfree"] | "quantale" => ["add", "Quantale"] + | "square" => ["even"] | x => [x] /-- @@ -1007,8 +1008,10 @@ def fixAbbreviation : List String → List String | "Comm" :: "Add" :: s => "AddComm" :: fixAbbreviation s | "Zero" :: "LE" :: s => "Nonneg" :: fixAbbreviation s | "zero" :: "_" :: "le" :: s => "nonneg" :: fixAbbreviation s + | "zero" :: "LE" :: s => "nonneg" :: fixAbbreviation s | "Zero" :: "LT" :: s => "Pos" :: fixAbbreviation s | "zero" :: "_" :: "lt" :: s => "pos" :: fixAbbreviation s + | "zero" :: "LT" :: s => "pos" :: fixAbbreviation s | "LE" :: "Zero" :: s => "Nonpos" :: fixAbbreviation s | "le" :: "_" :: "zero" :: s => "nonpos" :: fixAbbreviation s | "LT" :: "Zero" :: s => "Neg" :: fixAbbreviation s @@ -1025,9 +1028,8 @@ def fixAbbreviation : List String → List String | "Add" :: "Indicator" :: s => "Indicator" :: fixAbbreviation s | "add" :: "Indicator" :: s => "indicator" :: fixAbbreviation s | "add" :: "_" :: "indicator" :: s => "indicator" :: fixAbbreviation s - | "is" :: "Square" :: s => "even" :: fixAbbreviation s - | "Is" :: "Square" :: s => "Even" :: fixAbbreviation s - | "square" :: "In" :: s => "evenIn" :: fixAbbreviation s + | "is" :: "Even" :: s => "even" :: fixAbbreviation s + | "Is" :: "Even" :: s => "Even" :: fixAbbreviation s -- "Regular" is well-used in mathlib with various meanings (e.g. in -- measure theory) and a direct translation -- "regular" --> ["add", "Regular"] in `nameDict` above seems error-prone. diff --git a/Mathlib/Topology/Category/TopCat/Opens.lean b/Mathlib/Topology/Category/TopCat/Opens.lean index a34c444596d5a..3a0f0d9137b1e 100644 --- a/Mathlib/Topology/Category/TopCat/Opens.lean +++ b/Mathlib/Topology/Category/TopCat/Opens.lean @@ -409,7 +409,7 @@ lemma set_range_inclusion' {X : TopCat} (U : Opens X) : theorem functor_map_eq_inf {X : TopCat} (U V : Opens X) : U.isOpenEmbedding.isOpenMap.functor.obj ((Opens.map U.inclusion').obj V) = V ⊓ U := by ext1 - simp only [IsOpenMap.functor_obj_coe, map_coe, coe_inf, + simp only [IsOpenMap.coe_functor_obj, map_coe, coe_inf, Set.image_preimage_eq_inter_range, set_range_inclusion' U] theorem map_functor_eq' {X U : TopCat} (f : U ⟶ X) (hf : IsOpenEmbedding f) (V) : diff --git a/Mathlib/Topology/Sets/Closeds.lean b/Mathlib/Topology/Sets/Closeds.lean index 093d5981b92e7..f4e11a7bf12e2 100644 --- a/Mathlib/Topology/Sets/Closeds.lean +++ b/Mathlib/Topology/Sets/Closeds.lean @@ -274,7 +274,7 @@ theorem isClopen (s : Clopens α) : IsClopen (s : Set α) := /-- See Note [custom simps projection]. -/ def Simps.coe (s : Clopens α) : Set α := s -initialize_simps_projections Clopens (carrier → coe) +initialize_simps_projections Clopens (carrier → coe, as_prefix coe) /-- Reinterpret a clopen as an open. -/ @[simps] diff --git a/Mathlib/Topology/Sets/Compacts.lean b/Mathlib/Topology/Sets/Compacts.lean index 4cbbe40342a9b..fea42e6322ded 100644 --- a/Mathlib/Topology/Sets/Compacts.lean +++ b/Mathlib/Topology/Sets/Compacts.lean @@ -45,7 +45,7 @@ instance : SetLike (Compacts α) α where /-- See Note [custom simps projection]. -/ def Simps.coe (s : Compacts α) : Set α := s -initialize_simps_projections Compacts (carrier → coe) +initialize_simps_projections Compacts (carrier → coe, as_prefix coe) protected theorem isCompact (s : Compacts α) : IsCompact (s : Set α) := s.isCompact' @@ -197,7 +197,7 @@ instance : SetLike (NonemptyCompacts α) α where /-- See Note [custom simps projection]. -/ def Simps.coe (s : NonemptyCompacts α) : Set α := s -initialize_simps_projections NonemptyCompacts (carrier → coe) +initialize_simps_projections NonemptyCompacts (carrier → coe, as_prefix coe) protected theorem isCompact (s : NonemptyCompacts α) : IsCompact (s : Set α) := s.isCompact' @@ -288,7 +288,7 @@ instance : SetLike (PositiveCompacts α) α where /-- See Note [custom simps projection]. -/ def Simps.coe (s : PositiveCompacts α) : Set α := s -initialize_simps_projections PositiveCompacts (carrier → coe) +initialize_simps_projections PositiveCompacts (carrier → coe, as_prefix coe) protected theorem isCompact (s : PositiveCompacts α) : IsCompact (s : Set α) := s.isCompact' @@ -417,7 +417,7 @@ instance : SetLike (CompactOpens α) α where /-- See Note [custom simps projection]. -/ def Simps.coe (s : CompactOpens α) : Set α := s -initialize_simps_projections CompactOpens (carrier → coe) +initialize_simps_projections CompactOpens (carrier → coe, as_prefix coe) protected theorem isCompact (s : CompactOpens α) : IsCompact (s : Set α) := s.isCompact' diff --git a/Mathlib/Topology/Sets/Opens.lean b/Mathlib/Topology/Sets/Opens.lean index 452c6d12ef50e..955325b19092c 100644 --- a/Mathlib/Topology/Sets/Opens.lean +++ b/Mathlib/Topology/Sets/Opens.lean @@ -118,7 +118,7 @@ protected theorem isOpen (U : Opens α) : IsOpen (U : Set α) := /-- See Note [custom simps projection]. -/ def Simps.coe (U : Opens α) : Set α := U -initialize_simps_projections Opens (carrier → coe) +initialize_simps_projections Opens (carrier → coe, as_prefix coe) /-- The interior of a set, as an element of `Opens`. -/ @[simps] diff --git a/Mathlib/Topology/Sets/Order.lean b/Mathlib/Topology/Sets/Order.lean index c5cbb54e76bcb..56e0bc4255853 100644 --- a/Mathlib/Topology/Sets/Order.lean +++ b/Mathlib/Topology/Sets/Order.lean @@ -36,7 +36,7 @@ instance : SetLike (ClopenUpperSet α) α where /-- See Note [custom simps projection]. -/ def Simps.coe (s : ClopenUpperSet α) : Set α := s -initialize_simps_projections ClopenUpperSet (carrier → coe) +initialize_simps_projections ClopenUpperSet (carrier → coe, as_prefix coe) theorem upper (s : ClopenUpperSet α) : IsUpperSet (s : Set α) := s.upper' From 9b434d1b7fdce49c076c8c00ef5d57a4749233e6 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Fri, 3 Jan 2025 23:29:26 +0000 Subject: [PATCH 127/189] fix: make `Prod` projection delaborators respond to options, add option to disable (#20455) Adds option `pp.numericProj.prod` to configure whether or not to use `x.1` and `x.2` rather than `x.fst` and `x.snd`. Fixes the delaborators to make sure they respond to `pp.fieldNotation` and `pp.explicit`. Co-authored-by: Patrick Massot --- Mathlib/Data/Prod/Basic.lean | 40 +++++++++++++++++++++++++++-------- MathlibTest/delaborators.lean | 29 +++++++++++++++++++++++++ 2 files changed, 60 insertions(+), 9 deletions(-) diff --git a/Mathlib/Data/Prod/Basic.lean b/Mathlib/Data/Prod/Basic.lean index 7df6836ea3cac..37cb24c949727 100644 --- a/Mathlib/Data/Prod/Basic.lean +++ b/Mathlib/Data/Prod/Basic.lean @@ -300,21 +300,43 @@ theorem map_involutive [Nonempty α] [Nonempty β] {f : α → α} {g : β → Involutive (map f g) ↔ Involutive f ∧ Involutive g := map_leftInverse -section delaborators +namespace PrettyPrinting open Lean PrettyPrinter Delaborator +/-- +When true, then `Prod.fst x` and `Prod.snd x` pretty print as `x.1` and `x.2` +rather than as `x.fst` and `x.snd`. +-/ +register_option pp.numericProj.prod : Bool := { + defValue := true + descr := "enable pretty printing `Prod.fst x` as `x.1` and `Prod.snd x` as `x.2`." +} + +/-- Tell whether pretty-printing should use numeric projection notations `.1` +and `.2` for `Prod.fst` and `Prod.snd`. -/ +def getPPNumericProjProd (o : Options) : Bool := + o.get pp.numericProj.prod.name pp.numericProj.prod.defValue + /-- Delaborator for `Prod.fst x` as `x.1`. -/ @[app_delab Prod.fst] -def delabProdFst : Delab := withOverApp 3 do - let x ← SubExpr.withAppArg delab - `($(x).1) +def delabProdFst : Delab := + whenPPOption getPPNumericProjProd <| + whenPPOption getPPFieldNotation <| + whenNotPPOption getPPExplicit <| + withOverApp 3 do + let x ← SubExpr.withAppArg delab + `($(x).1) /-- Delaborator for `Prod.snd x` as `x.2`. -/ @[app_delab Prod.snd] -def delabProdSnd : Delab := withOverApp 3 do - let x ← SubExpr.withAppArg delab - `($(x).2) - -end delaborators +def delabProdSnd : Delab := + whenPPOption getPPNumericProjProd <| + whenPPOption getPPFieldNotation <| + whenNotPPOption getPPExplicit <| + withOverApp 3 do + let x ← SubExpr.withAppArg delab + `($(x).2) + +end PrettyPrinting end Prod diff --git a/MathlibTest/delaborators.lean b/MathlibTest/delaborators.lean index e29d29bfd0234..a0a0aaf448875 100644 --- a/MathlibTest/delaborators.lean +++ b/MathlibTest/delaborators.lean @@ -232,4 +232,33 @@ variable (p : (ℕ → ℕ) × (ℕ → ℕ)) #guard_msgs in #check p.1 22 +set_option pp.numericProj.prod false in +/-- info: x.fst : ℕ -/ +#guard_msgs in +#check x.1 + +set_option pp.numericProj.prod false in +/-- info: x.snd : ℕ -/ +#guard_msgs in +#check x.2 + +set_option pp.explicit true in +/-- info: @Prod.fst Nat Nat x : Nat -/ +#guard_msgs in +#check x.1 + +set_option pp.explicit true in +/-- info: @Prod.snd Nat Nat x : Nat -/ +#guard_msgs in +#check x.2 + +set_option pp.fieldNotation false in +/-- info: Prod.fst x : ℕ -/ +#guard_msgs in +#check x.1 + +set_option pp.fieldNotation false in +/-- info: Prod.snd x : ℕ -/ +#guard_msgs in +#check x.2 end prod From ae96bfcfb2ddccfa405843bd7ab6137c1082fde1 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Sat, 4 Jan 2025 04:52:50 +0000 Subject: [PATCH 128/189] chore: update Mathlib dependencies 2025-01-04 (#20463) This PR updates the Mathlib dependencies. --- lake-manifest.json | 4 ++-- lean-toolchain | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index 87602c47da393..d2f50138b644a 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -45,7 +45,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "a4a08d92be3de00def5298059bf707c72dfd3c66", + "rev": "2689851f387bb2cef351e6825fe94a56a304ca13", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "f007bfe46ea8fb801ec907df9ab540054abcc5fd", + "rev": "e8dc5fc16c625fc4fe08f42d625523275ddbbb4b", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", diff --git a/lean-toolchain b/lean-toolchain index cf25a9816f061..f8e93dad4c73c 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.15.0-rc1 +leanprover/lean4:v4.15.0 \ No newline at end of file From 29f9a66d622d9bab7f419120e22bb0d2598676ab Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 4 Jan 2025 05:18:40 +0000 Subject: [PATCH 129/189] chore: bump toolchain to v4.15.0 (#20461) --- lake-manifest.json | 18 +++++++++--------- lakefile.lean | 10 +++++----- lean-toolchain | 2 +- 3 files changed, 15 insertions(+), 15 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index d2f50138b644a..b6975fa1f851e 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,10 +5,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "8e5cb8d424df462f84997dd68af6f40e347c3e35", + "rev": "2c57364ef83406ea86d0f78ce3e342079a2fece5", "name": "plausible", "manifestFile": "lake-manifest.json", - "inputRev": "v4.15.0-rc1", + "inputRev": "v4.15.0", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/LeanSearchClient", @@ -25,10 +25,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "ed3b856bd8893ade75cafe13e8544d4c2660f377", + "rev": "9a0b533c2fbd6195df067630be18e11e4349051c", "name": "importGraph", "manifestFile": "lake-manifest.json", - "inputRev": "v4.15.0-rc1", + "inputRev": "v4.15.0", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/ProofWidgets4", @@ -48,19 +48,19 @@ "rev": "2689851f387bb2cef351e6825fe94a56a304ca13", "name": "aesop", "manifestFile": "lake-manifest.json", - "inputRev": "master", + "inputRev": "v4.15.0", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/quote4", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "ad942fdf0b15c38bface6acbb01d63855a2519ac", + "rev": "f0c584bcb14c5adfb53079781eeea75b26ebbd32", "name": "Qq", "manifestFile": "lake-manifest.json", - "inputRev": "v4.14.0", + "inputRev": "v4.15.0", "inherited": false, - "configFile": "lakefile.lean"}, + "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/batteries", "type": "git", "subDir": null, @@ -68,7 +68,7 @@ "rev": "e8dc5fc16c625fc4fe08f42d625523275ddbbb4b", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "main", + "inputRev": "v4.15.0", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lakefile.lean b/lakefile.lean index 1bf2f20f098ea..44bbb99b71251 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -7,13 +7,13 @@ open Lake DSL ## Mathlib dependencies on upstream projects -/ -require "leanprover-community" / "batteries" @ git "main" -require "leanprover-community" / "Qq" @ git "v4.14.0" -require "leanprover-community" / "aesop" @ git "master" +require "leanprover-community" / "batteries" @ git "v4.15.0" +require "leanprover-community" / "Qq" @ git "v4.15.0" +require "leanprover-community" / "aesop" @ git "v4.15.0" require "leanprover-community" / "proofwidgets" @ git "v0.0.48" -require "leanprover-community" / "importGraph" @ git "v4.15.0-rc1" +require "leanprover-community" / "importGraph" @ git "v4.15.0" require "leanprover-community" / "LeanSearchClient" @ git "main" -require "leanprover-community" / "plausible" @ git "v4.15.0-rc1" +require "leanprover-community" / "plausible" @ git "v4.15.0" /-! ## Options for building mathlib diff --git a/lean-toolchain b/lean-toolchain index f8e93dad4c73c..d0eb99ff68d76 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.15.0 \ No newline at end of file +leanprover/lean4:v4.15.0 From e71f6e9769283aa927ec7842e7e75e45f6f30324 Mon Sep 17 00:00:00 2001 From: Daniel Weber Date: Sat, 4 Jan 2025 06:32:00 +0000 Subject: [PATCH 130/189] feat(Algebra/Order/AddGroupWithTop): lemmas about LinearOrderedAddCommGroupWithTop (#18954) Co-authored-by: Daniel Weber <55664973+Command-Master@users.noreply.github.com> --- Mathlib/Algebra/Order/AddGroupWithTop.lean | 34 ++++++++++++++++++++++ 1 file changed, 34 insertions(+) diff --git a/Mathlib/Algebra/Order/AddGroupWithTop.lean b/Mathlib/Algebra/Order/AddGroupWithTop.lean index bb69b9cea49ec..909f38d98e085 100644 --- a/Mathlib/Algebra/Order/AddGroupWithTop.lean +++ b/Mathlib/Algebra/Order/AddGroupWithTop.lean @@ -194,4 +194,38 @@ instance (priority := 100) toSubtractionMonoid : SubtractionMonoid α where intro v simp [v] at h +lemma injective_add_left_of_ne_top (b : α) (h : b ≠ ⊤) : Function.Injective (fun x ↦ x + b) := by + intro x y h2 + replace h2 : x + (b + -b) = y + (b + -b) := by simp [← add_assoc, h2] + simpa only [LinearOrderedAddCommGroupWithTop.add_neg_cancel _ h, add_zero] using h2 + +lemma injective_add_right_of_ne_top (b : α) (h : b ≠ ⊤) : Function.Injective (fun x ↦ b + x) := by + simpa [add_comm] using injective_add_left_of_ne_top b h + +lemma strictMono_add_left_of_ne_top (b : α) (h : b ≠ ⊤) : StrictMono (fun x ↦ x + b) := by + apply Monotone.strictMono_of_injective + · apply Monotone.add_const monotone_id + · apply injective_add_left_of_ne_top _ h + +lemma strictMono_add_right_of_ne_top (b : α) (h : b ≠ ⊤) : StrictMono (fun x ↦ b + x) := by + simpa [add_comm] using strictMono_add_left_of_ne_top b h + +lemma sub_pos (a b : α) : 0 < a - b ↔ b < a ∨ b = ⊤ where + mp h := by + refine or_iff_not_imp_right.mpr fun h2 ↦ ?_ + replace h := strictMono_add_left_of_ne_top _ h2 h + simp only [zero_add] at h + rw [sub_eq_add_neg, add_assoc, add_comm (-b), + add_neg_cancel_of_ne_top h2, add_zero] at h + exact h + mpr h := by + rcases h with h | h + · convert strictMono_add_left_of_ne_top (-b) (by simp [h.ne_top]) h using 1 + · simp [add_neg_cancel_of_ne_top h.ne_top] + · simp [sub_eq_add_neg] + · rw [h] + simp only [sub_eq_add_neg, LinearOrderedAddCommGroupWithTop.neg_top, add_top] + apply lt_of_le_of_ne le_top + exact Ne.symm top_ne_zero + end LinearOrderedAddCommGroupWithTop From a97b21ff86ea58a827f0d95be7ff3d799647b417 Mon Sep 17 00:00:00 2001 From: Vasily Ilin Date: Sat, 4 Jan 2025 08:30:22 +0000 Subject: [PATCH 131/189] feat(Probability/Moments): add lemmas about moment generating functions (#19886) Add `mgf_ident_distrib`, `mgf_sum_iid`. These are some of the supporting lemmas we have proved on our way to formalizing the central limit theorem. Co-authored-by: Siyuan Ge --- Mathlib/Probability/Moments.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/Mathlib/Probability/Moments.lean b/Mathlib/Probability/Moments.lean index e6ac3e9af4f90..b6c17d414699d 100644 --- a/Mathlib/Probability/Moments.lean +++ b/Mathlib/Probability/Moments.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ +import Mathlib.Probability.IdentDistrib import Mathlib.Probability.Variance /-! @@ -275,6 +276,8 @@ theorem iIndepFun.integrable_exp_mul_sum [IsFiniteMeasure μ] {X : ι → Ω → refine IndepFun.integrable_exp_mul_add ?_ (h_int i (mem_insert_self _ _)) h_rec exact (h_indep.indepFun_finset_sum_of_not_mem h_meas hi_notin_s).symm +-- TODO(vilin97): weaken `h_meas` to `AEMeasurable (X i)` or `AEStronglyMeasurable (X i)` throughout +-- https://github.com/leanprover-community/mathlib4/issues/20367 theorem iIndepFun.mgf_sum {X : ι → Ω → ℝ} (h_indep : iIndepFun (fun _ => inferInstance) X μ) (h_meas : ∀ i, Measurable (X i)) (s : Finset ι) : mgf (∑ i ∈ s, X i) μ t = ∏ i ∈ s, mgf (X i) μ t := by @@ -299,6 +302,22 @@ theorem iIndepFun.cgf_sum {X : ι → Ω → ℝ} · rw [h_indep.mgf_sum h_meas] · exact (mgf_pos (h_int j hj)).ne' +theorem mgf_congr_of_identDistrib + (X : Ω → ℝ) {Ω' : Type*} {m' : MeasurableSpace Ω'} {μ' : Measure Ω'} (X' : Ω' → ℝ) + (hident : IdentDistrib X X' μ μ') (t : ℝ) : + mgf X μ t = mgf X' μ' t := hident.comp (measurable_const_mul t).exp |>.integral_eq + +theorem mgf_sum_of_identDistrib + {X : ι → Ω → ℝ} + {s : Finset ι} {j : ι} + (h_meas : ∀ i, Measurable (X i)) + (h_indep : iIndepFun (fun _ => inferInstance) X μ) + (hident : ∀ i ∈ s, ∀ j ∈ s, IdentDistrib (X i) (X j) μ μ) + (hj : j ∈ s) (t : ℝ) : mgf (∑ i ∈ s, X i) μ t = mgf (X j) μ t ^ #s := by + rw [h_indep.mgf_sum h_meas] + exact Finset.prod_eq_pow_card fun i hi => + mgf_congr_of_identDistrib (X i) (X j) (hident i hi j hj) t + /-- **Chernoff bound** on the upper tail of a real random variable. -/ theorem measure_ge_le_exp_mul_mgf [IsFiniteMeasure μ] (ε : ℝ) (ht : 0 ≤ t) (h_int : Integrable (fun ω => exp (t * X ω)) μ) : From b56ba00ef238036292a7f7d65fec7006a1f9c841 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Miyahara=20K=C5=8D?= Date: Sat, 4 Jan 2025 09:33:07 +0000 Subject: [PATCH 132/189] fix(Mathlib.Tactic.CC.Datatypes): `cc` raises panic (#20422) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit [Zulip thread](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/cc.20bug/near/491678555) The following `cc` raises a panic: ```lean example (n k : ℤ) (hnk : n = 2 * k + 1) (hk : (2 * k + 1) * (2 * k + 1 + 1) = 2 * ((2 * k + 1) * (k + 1))) : n * (n + 1) = 2 * ((2 * k + 1) * (k + 1)) := by cc ``` This is because `args₂[i]!` is evaluated before `compare args₁.size args₂.size` in the following code: ```lean scoped instance : Ord ACApps where compare | .ofExpr a, .ofExpr b => compare a b | .ofExpr _, .apps _ _ => .lt | .apps _ _, .ofExpr _ => .gt | .apps op₁ args₁, .apps op₂ args₂ => compare op₁ op₂ |>.then <| compare args₁.size args₂.size |>.then <| Id.run do for i in [:args₁.size] do let o := compare args₁[i]! args₂[i]! if o != .eq then return o return .eq ``` Usually, if `o₁ ≠ .eq`, `o₂` in `o₁.then o₂` is never evaluated because `Ordering.then` is marked by `@[macro_inline]`, however this doesn't work in this example. I don't know how `@[macro_inline]` works, but this bug can be solved by using `args₂[i]'h` instead of `args₂[i]!`. Also, this new definition is added to `Mathlib.Data.Ordering.Basic`: ```lean /-- `o₁.dthen fun h => o₂(h)` is like `o₁.then o₂` but `o₂` is allowed to depend on `h : o₁ = .eq`. -/ @[macro_inline] def dthen : (o : Ordering) → (o = .eq → Ordering) → Ordering | .eq, f => f rfl | o, _ => o ``` This PR increases imports percentage in `Mathlib.Tactic.CC.Datatypes`, but it doesn't matter because it's only 4 files. --- Mathlib/Data/Ordering/Basic.lean | 7 +++++++ Mathlib/Data/Ordering/Lemmas.lean | 4 ++++ Mathlib/Tactic/CC/Datatypes.lean | 10 +++++----- MathlibTest/cc.lean | 9 +++++++++ 4 files changed, 25 insertions(+), 5 deletions(-) diff --git a/Mathlib/Data/Ordering/Basic.lean b/Mathlib/Data/Ordering/Basic.lean index 45c2f5778c1d4..a6ec608f3bc39 100644 --- a/Mathlib/Data/Ordering/Basic.lean +++ b/Mathlib/Data/Ordering/Basic.lean @@ -38,6 +38,13 @@ def Compares [LT α] : Ordering → α → α → Prop @[simp] lemma compares_gt [LT α] (a b : α) : Compares gt a b = (a > b) := rfl +/-- `o₁.dthen fun h => o₂(h)` is like `o₁.then o₂` but `o₂` is allowed to depend on +`h : o₁ = .eq`. -/ +@[macro_inline] def dthen : + (o : Ordering) → (o = .eq → Ordering) → Ordering + | .eq, f => f rfl + | o, _ => o + end Ordering /-- diff --git a/Mathlib/Data/Ordering/Lemmas.lean b/Mathlib/Data/Ordering/Lemmas.lean index 42da3c30c0fb9..a5027ef17ad7c 100644 --- a/Mathlib/Data/Ordering/Lemmas.lean +++ b/Mathlib/Data/Ordering/Lemmas.lean @@ -29,6 +29,10 @@ theorem ite_eq_gt_distrib (c : Prop) [Decidable c] (a b : Ordering) : ((if c then a else b) = Ordering.gt) = if c then a = Ordering.gt else b = Ordering.gt := by by_cases c <;> simp [*] +@[simp] +lemma dthen_eq_then (o₁ o₂ : Ordering) : o₁.dthen (fun _ => o₂) = o₁.then o₂ := by + cases o₁ <;> rfl + end Ordering section diff --git a/Mathlib/Tactic/CC/Datatypes.lean b/Mathlib/Tactic/CC/Datatypes.lean index 32239f6b361a3..d45f190596b92 100644 --- a/Mathlib/Tactic/CC/Datatypes.lean +++ b/Mathlib/Tactic/CC/Datatypes.lean @@ -3,11 +3,10 @@ Copyright (c) 2016 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Miyahara Kō -/ -import Lean.Meta.CongrTheorems import Lean.Meta.Tactic.Rfl import Batteries.Data.RBMap.Basic import Mathlib.Lean.Meta.Basic -import Std.Data.HashMap.Basic +import Mathlib.Data.Ordering.Basic /-! # Datatypes for `cc` @@ -154,9 +153,10 @@ scoped instance : Ord ACApps where | .ofExpr _, .apps _ _ => .lt | .apps _ _, .ofExpr _ => .gt | .apps op₁ args₁, .apps op₂ args₂ => - compare op₁ op₂ |>.then <| compare args₁.size args₂.size |>.then <| Id.run do - for i in [:args₁.size] do - let o := compare args₁[i]! args₂[i]! + compare op₁ op₂ |>.then <| compare args₁.size args₂.size |>.dthen fun hs => Id.run do + have hs := Batteries.BEqCmp.cmp_iff_eq.mp hs + for hi : i in [:args₁.size] do + have hi := hi.right; let o := compare (args₁[i]'hi) (args₂[i]'(hs ▸ hi)) if o != .eq then return o return .eq diff --git a/MathlibTest/cc.lean b/MathlibTest/cc.lean index 8fd7047024b7e..7196296010426 100644 --- a/MathlibTest/cc.lean +++ b/MathlibTest/cc.lean @@ -592,3 +592,12 @@ example : "Miyahara Kō" = "Miyahara Kō" := by cc end lit + +section CCPanic + +example (n k : ℤ) (hnk : n = 2 * k + 1) + (hk : (2 * k + 1) * (2 * k + 1 + 1) = 2 * ((2 * k + 1) * (k + 1))) : + n * (n + 1) = 2 * ((2 * k + 1) * (k + 1)) := by + cc + +end CCPanic From c8e9d14233399d0c195b8dfd478fd96d4364fb10 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Sat, 4 Jan 2025 10:16:07 +0000 Subject: [PATCH 133/189] feat(MeasureTheory/Group/Measure): add ContinuousMulEquiv.isHaarMeasure_map (#20469) From FLT. --- Mathlib/MeasureTheory/Group/Measure.lean | 16 ++++++++++++++-- 1 file changed, 14 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 1c5dbcc1d6ff3..0dc9999c175a6 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -7,6 +7,7 @@ import Mathlib.MeasureTheory.Group.Action import Mathlib.MeasureTheory.Measure.Prod import Mathlib.Topology.Algebra.Module.Equiv import Mathlib.Topology.ContinuousMap.CocompactMap +import Mathlib.Topology.Algebra.ContinuousMonoidHom /-! # Measures on Groups @@ -704,10 +705,11 @@ theorem isHaarMeasure_of_isCompact_nonempty_interior [TopologicalGroup G] [Borel toIsOpenPosMeasure := isOpenPosMeasure_of_mulLeftInvariant_of_compact K hK h } /-- The image of a Haar measure under a continuous surjective proper group homomorphism is again -a Haar measure. See also `MulEquiv.isHaarMeasure_map`. -/ +a Haar measure. See also `MulEquiv.isHaarMeasure_map` and `ContinuousMulEquiv.isHaarMeasure_map`. -/ @[to_additive "The image of an additive Haar measure under a continuous surjective proper additive group -homomorphism is again an additive Haar measure. See also `AddEquiv.isAddHaarMeasure_map`."] +homomorphism is again an additive Haar measure. See also `AddEquiv.isAddHaarMeasure_map`, +`ContinuousAddEquiv.isAddHaarMeasure_map` and `ContinuousLinearEquiv.isAddHaarMeasure_map`."] theorem isHaarMeasure_map [BorelSpace G] [TopologicalGroup G] {H : Type*} [Group H] [TopologicalSpace H] [MeasurableSpace H] [BorelSpace H] [TopologicalGroup H] (f : G →* H) (hf : Continuous f) (h_surj : Surjective f) @@ -769,6 +771,16 @@ nonrec theorem _root_.MulEquiv.isHaarMeasure_map [BorelSpace G] [TopologicalGrou -/ isHaarMeasure_map μ e.toMonoidHom he e.surjective f.isClosedEmbedding.tendsto_cocompact +/-- +A convenience wrapper for MeasureTheory.Measure.isHaarMeasure_map. +-/ +@[to_additive "A convenience wrapper for MeasureTheory.Measure.isAddHaarMeasure_map. +"] +instance _root_.ContinuousMulEquiv.isHaarMeasure_map [BorelSpace G] [TopologicalGroup G] {H : Type*} + [Group H] [TopologicalSpace H] [MeasurableSpace H] [BorelSpace H] + [TopologicalGroup H] (e : G ≃ₜ* H) : (μ.map e).IsHaarMeasure := + e.toMulEquiv.isHaarMeasure_map μ e.continuous e.symm.continuous + /-- A convenience wrapper for MeasureTheory.Measure.isAddHaarMeasure_map`. -/ instance _root_.ContinuousLinearEquiv.isAddHaarMeasure_map {E F R S : Type*} [Semiring R] [Semiring S] From 9309505cf506dfc88977c51cdd71290df3ace31a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Sat, 4 Jan 2025 12:41:44 +0000 Subject: [PATCH 134/189] chore: split Kernel/MeasurableIntegral (#20427) The file is split into two files, one about the Lebesgue integral and one about the Bochner integral. The new MeasurableLIntegral file has around 500 less imports than MeasurableIntegral since it does not import Lp spaces and Bochner integrals. --- Mathlib.lean | 1 + .../Probability/Kernel/Composition/Basic.lean | 3 +- .../Kernel/Composition/IntegralCompProd.lean | 2 +- .../Kernel/MeasurableIntegral.lean | 201 +---------------- .../Kernel/MeasurableLIntegral.lean | 212 ++++++++++++++++++ Mathlib/Probability/Kernel/WithDensity.lean | 2 +- 6 files changed, 223 insertions(+), 198 deletions(-) create mode 100644 Mathlib/Probability/Kernel/MeasurableLIntegral.lean diff --git a/Mathlib.lean b/Mathlib.lean index 04ae5f192f56b..5cb43e13478e5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4271,6 +4271,7 @@ import Mathlib.Probability.Kernel.Disintegration.Unique import Mathlib.Probability.Kernel.Integral import Mathlib.Probability.Kernel.Invariance import Mathlib.Probability.Kernel.MeasurableIntegral +import Mathlib.Probability.Kernel.MeasurableLIntegral import Mathlib.Probability.Kernel.Proper import Mathlib.Probability.Kernel.RadonNikodym import Mathlib.Probability.Kernel.WithDensity diff --git a/Mathlib/Probability/Kernel/Composition/Basic.lean b/Mathlib/Probability/Kernel/Composition/Basic.lean index fd06f343aa595..30d4f488d8861 100644 --- a/Mathlib/Probability/Kernel/Composition/Basic.lean +++ b/Mathlib/Probability/Kernel/Composition/Basic.lean @@ -3,7 +3,8 @@ 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.MeasurableIntegral +import Mathlib.MeasureTheory.Measure.Prod +import Mathlib.Probability.Kernel.MeasurableLIntegral /-! # Product and composition of kernels diff --git a/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean b/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean index f8c3da0ddb92e..0f7fef4ed3332 100644 --- a/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean +++ b/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.Composition.Basic -import Mathlib.MeasureTheory.Integral.SetIntegral +import Mathlib.Probability.Kernel.MeasurableIntegral /-! # Bochner integral of a function against the composition-product of two kernels diff --git a/Mathlib/Probability/Kernel/MeasurableIntegral.lean b/Mathlib/Probability/Kernel/MeasurableIntegral.lean index 7c6f7a6e6bd19..d3eb71066d3ed 100644 --- a/Mathlib/Probability/Kernel/MeasurableIntegral.lean +++ b/Mathlib/Probability/Kernel/MeasurableIntegral.lean @@ -4,19 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.MeasureTheory.Integral.DominatedConvergence -import Mathlib.Probability.Kernel.Basic +import Mathlib.Probability.Kernel.MeasurableLIntegral /-! # Measurability of the integral against a kernel -The Lebesgue integral of a measurable function against a kernel is measurable. The Bochner integral -is strongly measurable. +The Bochner integral of a strongly measurable function against a kernel is strongly measurable. ## Main statements -* `Measurable.lintegral_kernel_prod_right`: the function `a ↦ ∫⁻ b, f a b ∂(κ a)` is measurable, - for an s-finite kernel `κ : Kernel α β` and a function `f : α → β → ℝ≥0∞` such that `uncurry f` - is measurable. * `MeasureTheory.StronglyMeasurable.integral_kernel_prod_right`: the function `a ↦ ∫ b, f a b ∂(κ a)` is measurable, for an s-finite kernel `κ : Kernel α β` and a function `f : α → β → E` such that `uncurry f` is measurable. @@ -30,204 +26,19 @@ open scoped MeasureTheory ENNReal Topology variable {α β γ : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} {κ : Kernel α β} {η : Kernel (α × β) γ} {a : α} + {E : Type*} [NormedAddCommGroup E] [IsSFiniteKernel κ] [IsSFiniteKernel η] -namespace ProbabilityTheory - -namespace Kernel - -/-- This is an auxiliary lemma for `measurable_kernel_prod_mk_left`. -/ -theorem measurable_kernel_prod_mk_left_of_finite {t : Set (α × β)} (ht : MeasurableSet t) - (hκs : ∀ a, IsFiniteMeasure (κ a)) : Measurable fun a => κ a (Prod.mk a ⁻¹' t) := by - -- `t` is a measurable set in the product `α × β`: we use that the product σ-algebra is generated - -- by boxes to prove the result by induction. - induction t, ht - using MeasurableSpace.induction_on_inter generateFrom_prod.symm isPiSystem_prod with - | empty => - simp only [preimage_empty, measure_empty, measurable_const] - | basic t ht => - simp only [Set.mem_image2, Set.mem_setOf_eq] at ht - obtain ⟨t₁, ht₁, t₂, ht₂, rfl⟩ := ht - classical - simp_rw [mk_preimage_prod_right_eq_if] - have h_eq_ite : (fun a => κ a (ite (a ∈ t₁) t₂ ∅)) = fun a => ite (a ∈ t₁) (κ a t₂) 0 := by - ext1 a - split_ifs - exacts [rfl, measure_empty] - rw [h_eq_ite] - exact Measurable.ite ht₁ (Kernel.measurable_coe κ ht₂) measurable_const - | compl t htm iht => - have h_eq_sdiff : ∀ a, Prod.mk a ⁻¹' tᶜ = Set.univ \ Prod.mk a ⁻¹' t := by - intro a - ext1 b - simp only [mem_compl_iff, mem_preimage, mem_diff, mem_univ, true_and] - simp_rw [h_eq_sdiff] - have : - (fun a => κ a (Set.univ \ Prod.mk a ⁻¹' t)) = fun a => - κ a Set.univ - κ a (Prod.mk a ⁻¹' t) := by - ext1 a - rw [← Set.diff_inter_self_eq_diff, Set.inter_univ, measure_diff (Set.subset_univ _)] - · exact (measurable_prod_mk_left htm).nullMeasurableSet - · exact measure_ne_top _ _ - rw [this] - exact Measurable.sub (Kernel.measurable_coe κ MeasurableSet.univ) iht - | iUnion f h_disj hf_meas hf => - have (a : α) : κ a (Prod.mk a ⁻¹' ⋃ i, f i) = ∑' i, κ a (Prod.mk a ⁻¹' f i) := by - rw [preimage_iUnion, measure_iUnion] - · exact h_disj.mono fun _ _ ↦ .preimage _ - · exact fun i ↦ measurable_prod_mk_left (hf_meas i) - simpa only [this] using Measurable.ennreal_tsum hf - -theorem measurable_kernel_prod_mk_left [IsSFiniteKernel κ] {t : Set (α × β)} - (ht : MeasurableSet t) : Measurable fun a => κ a (Prod.mk a ⁻¹' t) := by - rw [← Kernel.kernel_sum_seq κ] - have : ∀ a, Kernel.sum (Kernel.seq κ) a (Prod.mk a ⁻¹' t) = - ∑' n, Kernel.seq κ n a (Prod.mk a ⁻¹' t) := fun a => - Kernel.sum_apply' _ _ (measurable_prod_mk_left ht) - simp_rw [this] - refine Measurable.ennreal_tsum fun n => ?_ - exact measurable_kernel_prod_mk_left_of_finite ht inferInstance - -theorem measurable_kernel_prod_mk_left' [IsSFiniteKernel η] {s : Set (β × γ)} (hs : MeasurableSet s) - (a : α) : Measurable fun b => η (a, b) (Prod.mk b ⁻¹' s) := by - have : ∀ b, Prod.mk b ⁻¹' s = {c | ((a, b), c) ∈ {p : (α × β) × γ | (p.1.2, p.2) ∈ s}} := by - intro b; rfl - simp_rw [this] - refine (measurable_kernel_prod_mk_left ?_).comp measurable_prod_mk_left - exact (measurable_fst.snd.prod_mk measurable_snd) hs - -theorem measurable_kernel_prod_mk_right [IsSFiniteKernel κ] {s : Set (β × α)} - (hs : MeasurableSet s) : Measurable fun y => κ y ((fun x => (x, y)) ⁻¹' s) := - measurable_kernel_prod_mk_left (measurableSet_swap_iff.mpr hs) - -end Kernel - -open ProbabilityTheory.Kernel - -section Lintegral - -variable [IsSFiniteKernel κ] [IsSFiniteKernel η] - -/-- Auxiliary lemma for `Measurable.lintegral_kernel_prod_right`. -/ -theorem Kernel.measurable_lintegral_indicator_const {t : Set (α × β)} (ht : MeasurableSet t) - (c : ℝ≥0∞) : Measurable fun a => ∫⁻ b, t.indicator (Function.const (α × β) c) (a, b) ∂κ a := by - -- Porting note: was originally by - -- `simp_rw [lintegral_indicator_const_comp measurable_prod_mk_left ht _]` - -- but this has no effect, so added the `conv` below - conv => - congr - ext - erw [lintegral_indicator_const_comp measurable_prod_mk_left ht _] - exact Measurable.const_mul (measurable_kernel_prod_mk_left ht) c - -/-- For an s-finite kernel `κ` and a function `f : α → β → ℝ≥0∞` which is measurable when seen as a -map from `α × β` (hypothesis `Measurable (uncurry f)`), the integral `a ↦ ∫⁻ b, f a b ∂(κ a)` is -measurable. -/ -theorem _root_.Measurable.lintegral_kernel_prod_right {f : α → β → ℝ≥0∞} - (hf : Measurable (uncurry f)) : Measurable fun a => ∫⁻ b, f a b ∂κ a := by - let F : ℕ → SimpleFunc (α × β) ℝ≥0∞ := SimpleFunc.eapprox (uncurry f) - have h : ∀ a, ⨆ n, F n a = uncurry f a := SimpleFunc.iSup_eapprox_apply hf - simp only [Prod.forall, uncurry_apply_pair] at h - simp_rw [← h] - have : ∀ a, (∫⁻ b, ⨆ n, F n (a, b) ∂κ a) = ⨆ n, ∫⁻ b, F n (a, b) ∂κ a := by - intro a - rw [lintegral_iSup] - · exact fun n => (F n).measurable.comp measurable_prod_mk_left - · exact fun i j hij b => SimpleFunc.monotone_eapprox (uncurry f) hij _ - simp_rw [this] - refine .iSup fun n => ?_ - refine SimpleFunc.induction - (P := fun f => Measurable (fun (a : α) => ∫⁻ (b : β), f (a, b) ∂κ a)) ?_ ?_ (F n) - · intro c t ht - simp only [SimpleFunc.const_zero, SimpleFunc.coe_piecewise, SimpleFunc.coe_const, - SimpleFunc.coe_zero, Set.piecewise_eq_indicator] - exact Kernel.measurable_lintegral_indicator_const (κ := κ) ht c - · intro g₁ g₂ _ hm₁ hm₂ - simp only [SimpleFunc.coe_add, Pi.add_apply] - have h_add : - (fun a => ∫⁻ b, g₁ (a, b) + g₂ (a, b) ∂κ a) = - (fun a => ∫⁻ b, g₁ (a, b) ∂κ a) + fun a => ∫⁻ b, g₂ (a, b) ∂κ a := by - ext1 a - rw [Pi.add_apply] - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11224): was `rw` (`Function.comp` reducibility) - erw [lintegral_add_left (g₁.measurable.comp measurable_prod_mk_left)] - simp_rw [Function.comp_apply] - rw [h_add] - exact Measurable.add hm₁ hm₂ - -theorem _root_.Measurable.lintegral_kernel_prod_right' {f : α × β → ℝ≥0∞} (hf : Measurable f) : - Measurable fun a => ∫⁻ b, f (a, b) ∂κ a := by - refine Measurable.lintegral_kernel_prod_right ?_ - have : (uncurry fun (a : α) (b : β) => f (a, b)) = f := by - ext x; rw [uncurry_apply_pair] - rwa [this] - -theorem _root_.Measurable.lintegral_kernel_prod_right'' {f : β × γ → ℝ≥0∞} (hf : Measurable f) : - Measurable fun x => ∫⁻ y, f (x, y) ∂η (a, x) := by - -- Porting note: used `Prod.mk a` instead of `fun x => (a, x)` below - change - Measurable - ((fun x => ∫⁻ y, (fun u : (α × β) × γ => f (u.1.2, u.2)) (x, y) ∂η x) ∘ Prod.mk a) - -- Porting note: specified `κ`, `f`. - refine (Measurable.lintegral_kernel_prod_right' (κ := η) - (f := (fun u ↦ f (u.fst.snd, u.snd))) ?_).comp measurable_prod_mk_left - exact hf.comp (measurable_fst.snd.prod_mk measurable_snd) - -theorem _root_.Measurable.setLIntegral_kernel_prod_right {f : α → β → ℝ≥0∞} - (hf : Measurable (uncurry f)) {s : Set β} (hs : MeasurableSet s) : - Measurable fun a => ∫⁻ b in s, f a b ∂κ a := by - simp_rw [← lintegral_restrict κ hs]; exact hf.lintegral_kernel_prod_right - -@[deprecated (since := "2024-06-29")] -alias _root_.Measurable.set_lintegral_kernel_prod_right := - _root_.Measurable.setLIntegral_kernel_prod_right - -theorem _root_.Measurable.lintegral_kernel_prod_left' {f : β × α → ℝ≥0∞} (hf : Measurable f) : - Measurable fun y => ∫⁻ x, f (x, y) ∂κ y := - (measurable_swap_iff.mpr hf).lintegral_kernel_prod_right' - -theorem _root_.Measurable.lintegral_kernel_prod_left {f : β → α → ℝ≥0∞} - (hf : Measurable (uncurry f)) : Measurable fun y => ∫⁻ x, f x y ∂κ y := - hf.lintegral_kernel_prod_left' - -theorem _root_.Measurable.setLIntegral_kernel_prod_left {f : β → α → ℝ≥0∞} - (hf : Measurable (uncurry f)) {s : Set β} (hs : MeasurableSet s) : - Measurable fun b => ∫⁻ a in s, f a b ∂κ b := by - simp_rw [← lintegral_restrict κ hs]; exact hf.lintegral_kernel_prod_left - -@[deprecated (since := "2024-06-29")] -alias _root_.Measurable.set_lintegral_kernel_prod_left := - _root_.Measurable.setLIntegral_kernel_prod_left - -theorem _root_.Measurable.lintegral_kernel {f : β → ℝ≥0∞} (hf : Measurable f) : - Measurable fun a => ∫⁻ b, f b ∂κ a := - Measurable.lintegral_kernel_prod_right (hf.comp measurable_snd) - -theorem _root_.Measurable.setLIntegral_kernel {f : β → ℝ≥0∞} (hf : Measurable f) {s : Set β} - (hs : MeasurableSet s) : Measurable fun a => ∫⁻ b in s, f b ∂κ a := by - -- Porting note: was term mode proof (`Function.comp` reducibility) - refine Measurable.setLIntegral_kernel_prod_right ?_ hs - convert hf.comp measurable_snd - -@[deprecated (since := "2024-06-29")] -alias _root_.Measurable.set_lintegral_kernel := _root_.Measurable.setLIntegral_kernel - -end Lintegral - -variable {E : Type*} [NormedAddCommGroup E] [IsSFiniteKernel κ] - -theorem measurableSet_kernel_integrable ⦃f : α → β → E⦄ (hf : StronglyMeasurable (uncurry f)) : +theorem ProbabilityTheory.measurableSet_kernel_integrable ⦃f : α → β → E⦄ + (hf : StronglyMeasurable (uncurry f)) : MeasurableSet {x | Integrable (f x) (κ x)} := by simp_rw [Integrable, hf.of_uncurry_left.aestronglyMeasurable, true_and] exact measurableSet_lt (Measurable.lintegral_kernel_prod_right hf.ennnorm) measurable_const -end ProbabilityTheory - open ProbabilityTheory.Kernel namespace MeasureTheory -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [IsSFiniteKernel κ] - [IsSFiniteKernel η] +variable [NormedSpace ℝ E] theorem StronglyMeasurable.integral_kernel_prod_right ⦃f : α → β → E⦄ (hf : StronglyMeasurable (uncurry f)) : StronglyMeasurable fun x => ∫ y, f x y ∂κ x := by diff --git a/Mathlib/Probability/Kernel/MeasurableLIntegral.lean b/Mathlib/Probability/Kernel/MeasurableLIntegral.lean new file mode 100644 index 0000000000000..a5c878aab3b61 --- /dev/null +++ b/Mathlib/Probability/Kernel/MeasurableLIntegral.lean @@ -0,0 +1,212 @@ +/- +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.MeasureTheory.MeasurableSpace.Prod +import Mathlib.Probability.Kernel.Basic + +/-! +# Measurability of the integral against a kernel + +The Lebesgue integral of a measurable function against a kernel is measurable. + +## Main statements + +* `Measurable.lintegral_kernel_prod_right`: the function `a ↦ ∫⁻ b, f a b ∂(κ a)` is measurable, + for an s-finite kernel `κ : Kernel α β` and a function `f : α → β → ℝ≥0∞` such that `uncurry f` + is measurable. + +-/ + + +open MeasureTheory ProbabilityTheory Function Set Filter + +open scoped MeasureTheory ENNReal Topology + +variable {α β γ : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} + {κ : Kernel α β} {η : Kernel (α × β) γ} {a : α} + +namespace ProbabilityTheory + +namespace Kernel + +/-- This is an auxiliary lemma for `measurable_kernel_prod_mk_left`. -/ +theorem measurable_kernel_prod_mk_left_of_finite {t : Set (α × β)} (ht : MeasurableSet t) + (hκs : ∀ a, IsFiniteMeasure (κ a)) : Measurable fun a => κ a (Prod.mk a ⁻¹' t) := by + -- `t` is a measurable set in the product `α × β`: we use that the product σ-algebra is generated + -- by boxes to prove the result by induction. + induction t, ht + using MeasurableSpace.induction_on_inter generateFrom_prod.symm isPiSystem_prod with + | empty => + simp only [preimage_empty, measure_empty, measurable_const] + | basic t ht => + simp only [Set.mem_image2, Set.mem_setOf_eq] at ht + obtain ⟨t₁, ht₁, t₂, ht₂, rfl⟩ := ht + classical + simp_rw [mk_preimage_prod_right_eq_if] + have h_eq_ite : (fun a => κ a (ite (a ∈ t₁) t₂ ∅)) = fun a => ite (a ∈ t₁) (κ a t₂) 0 := by + ext1 a + split_ifs + exacts [rfl, measure_empty] + rw [h_eq_ite] + exact Measurable.ite ht₁ (Kernel.measurable_coe κ ht₂) measurable_const + | compl t htm iht => + have h_eq_sdiff : ∀ a, Prod.mk a ⁻¹' tᶜ = Set.univ \ Prod.mk a ⁻¹' t := by + intro a + ext1 b + simp only [mem_compl_iff, mem_preimage, mem_diff, mem_univ, true_and] + simp_rw [h_eq_sdiff] + have : + (fun a => κ a (Set.univ \ Prod.mk a ⁻¹' t)) = fun a => + κ a Set.univ - κ a (Prod.mk a ⁻¹' t) := by + ext1 a + rw [← Set.diff_inter_self_eq_diff, Set.inter_univ, measure_diff (Set.subset_univ _)] + · exact (measurable_prod_mk_left htm).nullMeasurableSet + · exact measure_ne_top _ _ + rw [this] + exact Measurable.sub (Kernel.measurable_coe κ MeasurableSet.univ) iht + | iUnion f h_disj hf_meas hf => + have (a : α) : κ a (Prod.mk a ⁻¹' ⋃ i, f i) = ∑' i, κ a (Prod.mk a ⁻¹' f i) := by + rw [preimage_iUnion, measure_iUnion] + · exact h_disj.mono fun _ _ ↦ .preimage _ + · exact fun i ↦ measurable_prod_mk_left (hf_meas i) + simpa only [this] using Measurable.ennreal_tsum hf + +theorem measurable_kernel_prod_mk_left [IsSFiniteKernel κ] {t : Set (α × β)} + (ht : MeasurableSet t) : Measurable fun a => κ a (Prod.mk a ⁻¹' t) := by + rw [← Kernel.kernel_sum_seq κ] + have : ∀ a, Kernel.sum (Kernel.seq κ) a (Prod.mk a ⁻¹' t) = + ∑' n, Kernel.seq κ n a (Prod.mk a ⁻¹' t) := fun a => + Kernel.sum_apply' _ _ (measurable_prod_mk_left ht) + simp_rw [this] + refine Measurable.ennreal_tsum fun n => ?_ + exact measurable_kernel_prod_mk_left_of_finite ht inferInstance + +theorem measurable_kernel_prod_mk_left' [IsSFiniteKernel η] {s : Set (β × γ)} (hs : MeasurableSet s) + (a : α) : Measurable fun b => η (a, b) (Prod.mk b ⁻¹' s) := by + have : ∀ b, Prod.mk b ⁻¹' s = {c | ((a, b), c) ∈ {p : (α × β) × γ | (p.1.2, p.2) ∈ s}} := by + intro b; rfl + simp_rw [this] + refine (measurable_kernel_prod_mk_left ?_).comp measurable_prod_mk_left + exact (measurable_fst.snd.prod_mk measurable_snd) hs + +theorem measurable_kernel_prod_mk_right [IsSFiniteKernel κ] {s : Set (β × α)} + (hs : MeasurableSet s) : Measurable fun y => κ y ((fun x => (x, y)) ⁻¹' s) := + measurable_kernel_prod_mk_left (measurableSet_swap_iff.mpr hs) + +end Kernel + +open ProbabilityTheory.Kernel + +section Lintegral + +variable [IsSFiniteKernel κ] [IsSFiniteKernel η] + +/-- Auxiliary lemma for `Measurable.lintegral_kernel_prod_right`. -/ +theorem Kernel.measurable_lintegral_indicator_const {t : Set (α × β)} (ht : MeasurableSet t) + (c : ℝ≥0∞) : Measurable fun a => ∫⁻ b, t.indicator (Function.const (α × β) c) (a, b) ∂κ a := by + -- Porting note: was originally by + -- `simp_rw [lintegral_indicator_const_comp measurable_prod_mk_left ht _]` + -- but this has no effect, so added the `conv` below + conv => + congr + ext + erw [lintegral_indicator_const_comp measurable_prod_mk_left ht _] + exact Measurable.const_mul (measurable_kernel_prod_mk_left ht) c + +/-- For an s-finite kernel `κ` and a function `f : α → β → ℝ≥0∞` which is measurable when seen as a +map from `α × β` (hypothesis `Measurable (uncurry f)`), the integral `a ↦ ∫⁻ b, f a b ∂(κ a)` is +measurable. -/ +theorem _root_.Measurable.lintegral_kernel_prod_right {f : α → β → ℝ≥0∞} + (hf : Measurable (uncurry f)) : Measurable fun a => ∫⁻ b, f a b ∂κ a := by + let F : ℕ → SimpleFunc (α × β) ℝ≥0∞ := SimpleFunc.eapprox (uncurry f) + have h : ∀ a, ⨆ n, F n a = uncurry f a := SimpleFunc.iSup_eapprox_apply hf + simp only [Prod.forall, uncurry_apply_pair] at h + simp_rw [← h] + have : ∀ a, (∫⁻ b, ⨆ n, F n (a, b) ∂κ a) = ⨆ n, ∫⁻ b, F n (a, b) ∂κ a := by + intro a + rw [lintegral_iSup] + · exact fun n => (F n).measurable.comp measurable_prod_mk_left + · exact fun i j hij b => SimpleFunc.monotone_eapprox (uncurry f) hij _ + simp_rw [this] + refine .iSup fun n => ?_ + refine SimpleFunc.induction + (P := fun f => Measurable (fun (a : α) => ∫⁻ (b : β), f (a, b) ∂κ a)) ?_ ?_ (F n) + · intro c t ht + simp only [SimpleFunc.const_zero, SimpleFunc.coe_piecewise, SimpleFunc.coe_const, + SimpleFunc.coe_zero, Set.piecewise_eq_indicator] + exact Kernel.measurable_lintegral_indicator_const (κ := κ) ht c + · intro g₁ g₂ _ hm₁ hm₂ + simp only [SimpleFunc.coe_add, Pi.add_apply] + have h_add : + (fun a => ∫⁻ b, g₁ (a, b) + g₂ (a, b) ∂κ a) = + (fun a => ∫⁻ b, g₁ (a, b) ∂κ a) + fun a => ∫⁻ b, g₂ (a, b) ∂κ a := by + ext1 a + rw [Pi.add_apply] + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11224): was `rw` (`Function.comp` reducibility) + erw [lintegral_add_left (g₁.measurable.comp measurable_prod_mk_left)] + simp_rw [Function.comp_apply] + rw [h_add] + exact Measurable.add hm₁ hm₂ + +theorem _root_.Measurable.lintegral_kernel_prod_right' {f : α × β → ℝ≥0∞} (hf : Measurable f) : + Measurable fun a => ∫⁻ b, f (a, b) ∂κ a := by + refine Measurable.lintegral_kernel_prod_right ?_ + have : (uncurry fun (a : α) (b : β) => f (a, b)) = f := by + ext x; rw [uncurry_apply_pair] + rwa [this] + +theorem _root_.Measurable.lintegral_kernel_prod_right'' {f : β × γ → ℝ≥0∞} (hf : Measurable f) : + Measurable fun x => ∫⁻ y, f (x, y) ∂η (a, x) := by + -- Porting note: used `Prod.mk a` instead of `fun x => (a, x)` below + change + Measurable + ((fun x => ∫⁻ y, (fun u : (α × β) × γ => f (u.1.2, u.2)) (x, y) ∂η x) ∘ Prod.mk a) + -- Porting note: specified `κ`, `f`. + refine (Measurable.lintegral_kernel_prod_right' (κ := η) + (f := (fun u ↦ f (u.fst.snd, u.snd))) ?_).comp measurable_prod_mk_left + exact hf.comp (measurable_fst.snd.prod_mk measurable_snd) + +theorem _root_.Measurable.setLIntegral_kernel_prod_right {f : α → β → ℝ≥0∞} + (hf : Measurable (uncurry f)) {s : Set β} (hs : MeasurableSet s) : + Measurable fun a => ∫⁻ b in s, f a b ∂κ a := by + simp_rw [← lintegral_restrict κ hs]; exact hf.lintegral_kernel_prod_right + +@[deprecated (since := "2024-06-29")] +alias _root_.Measurable.set_lintegral_kernel_prod_right := + _root_.Measurable.setLIntegral_kernel_prod_right + +theorem _root_.Measurable.lintegral_kernel_prod_left' {f : β × α → ℝ≥0∞} (hf : Measurable f) : + Measurable fun y => ∫⁻ x, f (x, y) ∂κ y := + (measurable_swap_iff.mpr hf).lintegral_kernel_prod_right' + +theorem _root_.Measurable.lintegral_kernel_prod_left {f : β → α → ℝ≥0∞} + (hf : Measurable (uncurry f)) : Measurable fun y => ∫⁻ x, f x y ∂κ y := + hf.lintegral_kernel_prod_left' + +theorem _root_.Measurable.setLIntegral_kernel_prod_left {f : β → α → ℝ≥0∞} + (hf : Measurable (uncurry f)) {s : Set β} (hs : MeasurableSet s) : + Measurable fun b => ∫⁻ a in s, f a b ∂κ b := by + simp_rw [← lintegral_restrict κ hs]; exact hf.lintegral_kernel_prod_left + +@[deprecated (since := "2024-06-29")] +alias _root_.Measurable.set_lintegral_kernel_prod_left := + _root_.Measurable.setLIntegral_kernel_prod_left + +theorem _root_.Measurable.lintegral_kernel {f : β → ℝ≥0∞} (hf : Measurable f) : + Measurable fun a => ∫⁻ b, f b ∂κ a := + Measurable.lintegral_kernel_prod_right (hf.comp measurable_snd) + +theorem _root_.Measurable.setLIntegral_kernel {f : β → ℝ≥0∞} (hf : Measurable f) {s : Set β} + (hs : MeasurableSet s) : Measurable fun a => ∫⁻ b in s, f b ∂κ a := by + -- Porting note: was term mode proof (`Function.comp` reducibility) + refine Measurable.setLIntegral_kernel_prod_right ?_ hs + convert hf.comp measurable_snd + +@[deprecated (since := "2024-06-29")] +alias _root_.Measurable.set_lintegral_kernel := _root_.Measurable.setLIntegral_kernel + +end Lintegral + +end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/WithDensity.lean b/Mathlib/Probability/Kernel/WithDensity.lean index a20305b92df90..172648c0563f8 100644 --- a/Mathlib/Probability/Kernel/WithDensity.lean +++ b/Mathlib/Probability/Kernel/WithDensity.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.MeasurableIntegral +import Mathlib.Probability.Kernel.MeasurableLIntegral import Mathlib.MeasureTheory.Integral.SetIntegral /-! From e6514e003e72760a7506240a95c3679ee828c8ab Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Sat, 4 Jan 2025 12:51:10 +0000 Subject: [PATCH 135/189] chore: split AEEqOfIntegral into two files, one for each integral type (#20405) The proof of `AEMeasurable.ae_eq_of_forall_setLIntegral_eq` is changed significantly since it used to call the same result for the Bochner integral. Everything else is copied as it was. --- Mathlib.lean | 1 + .../MeasureTheory/Decomposition/Lebesgue.lean | 3 +- .../Function/AEEqOfIntegral.lean | 183 +-------------- .../Function/AEEqOfLIntegral.lean | 210 ++++++++++++++++++ 4 files changed, 216 insertions(+), 181 deletions(-) create mode 100644 Mathlib/MeasureTheory/Function/AEEqOfLIntegral.lean diff --git a/Mathlib.lean b/Mathlib.lean index 5cb43e13478e5..4ebb3fc8f8431 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3661,6 +3661,7 @@ import Mathlib.MeasureTheory.Decomposition.UnsignedHahn import Mathlib.MeasureTheory.Function.AEEqFun import Mathlib.MeasureTheory.Function.AEEqFun.DomAct import Mathlib.MeasureTheory.Function.AEEqOfIntegral +import Mathlib.MeasureTheory.Function.AEEqOfLIntegral import Mathlib.MeasureTheory.Function.AEMeasurableOrder import Mathlib.MeasureTheory.Function.AEMeasurableSequence import Mathlib.MeasureTheory.Function.ConditionalExpectation.AEMeasurable diff --git a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean index f44f641788f35..f1cececd23a18 100644 --- a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ import Mathlib.MeasureTheory.Decomposition.UnsignedHahn -import Mathlib.MeasureTheory.Function.AEEqOfIntegral +import Mathlib.MeasureTheory.Function.AEEqOfLIntegral +import Mathlib.MeasureTheory.Function.L1Space import Mathlib.MeasureTheory.Measure.Sub /-! diff --git a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean index c6b139c016199..99e6748be5ef4 100644 --- a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean +++ b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean @@ -5,6 +5,7 @@ Authors: Rémy Degenne -/ import Mathlib.Analysis.InnerProductSpace.Basic import Mathlib.Analysis.Normed.Module.Dual +import Mathlib.MeasureTheory.Function.AEEqOfLIntegral import Mathlib.MeasureTheory.Function.StronglyMeasurable.Lp import Mathlib.MeasureTheory.Integral.SetIntegral import Mathlib.Order.Filter.Ring @@ -16,6 +17,8 @@ on all sets, then they are equal almost everywhere". The different lemmas use various hypotheses on the class of functions, on the target space or on the possible finiteness of the measure. +This file is about Bochner integrals. See the file `AEEqOfLIntegral` for Lebesgue integrals. + ## Main statements All results listed below apply to two functions `f, g`, together with two main hypotheses, @@ -31,9 +34,6 @@ The conclusion is then `f =ᵐ[μ] g`. The main lemmas are: For each of these results, we also provide a lemma about the equality of one function and 0. For example, `Lp.ae_eq_zero_of_forall_setIntegral_eq_zero`. -We also register the corresponding lemma for integrals of `ℝ≥0∞`-valued functions, in -`ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite`. - Generally useful lemmas which are not related to integrals: * `ae_eq_zero_of_forall_inner`: if for all constants `c`, `fun x => inner c (f x) =ᵐ[μ] 0` then `f =ᵐ[μ] 0`. @@ -118,132 +118,6 @@ variable {α E : Type*} {m m0 : MeasurableSpace α} {μ : Measure α} section AeEqOfForallSetIntegralEq -theorem ae_const_le_iff_forall_lt_measure_zero {β} [LinearOrder β] [TopologicalSpace β] - [OrderTopology β] [FirstCountableTopology β] (f : α → β) (c : β) : - (∀ᵐ x ∂μ, c ≤ f x) ↔ ∀ b < c, μ {x | f x ≤ b} = 0 := by - rw [ae_iff] - push_neg - constructor - · intro h b hb - exact measure_mono_null (fun y hy => (lt_of_le_of_lt hy hb : _)) h - intro hc - by_cases h : ∀ b, c ≤ b - · have : {a : α | f a < c} = ∅ := by - apply Set.eq_empty_iff_forall_not_mem.2 fun x hx => ?_ - exact (lt_irrefl _ (lt_of_lt_of_le hx (h (f x)))).elim - simp [this] - by_cases H : ¬IsLUB (Set.Iio c) c - · have : c ∈ upperBounds (Set.Iio c) := fun y hy => le_of_lt hy - obtain ⟨b, b_up, bc⟩ : ∃ b : β, b ∈ upperBounds (Set.Iio c) ∧ b < c := by - simpa [IsLUB, IsLeast, this, lowerBounds] using H - exact measure_mono_null (fun x hx => b_up hx) (hc b bc) - push_neg at H h - obtain ⟨u, _, u_lt, u_lim, -⟩ : - ∃ u : ℕ → β, - StrictMono u ∧ (∀ n : ℕ, u n < c) ∧ Tendsto u atTop (𝓝 c) ∧ ∀ n : ℕ, u n ∈ Set.Iio c := - H.exists_seq_strictMono_tendsto_of_not_mem (lt_irrefl c) h - have h_Union : {x | f x < c} = ⋃ n : ℕ, {x | f x ≤ u n} := by - ext1 x - simp_rw [Set.mem_iUnion, Set.mem_setOf_eq] - constructor <;> intro h - · obtain ⟨n, hn⟩ := ((tendsto_order.1 u_lim).1 _ h).exists; exact ⟨n, hn.le⟩ - · obtain ⟨n, hn⟩ := h; exact hn.trans_lt (u_lt _) - rw [h_Union, measure_iUnion_null_iff] - intro n - exact hc _ (u_lt n) - -section ENNReal - -open scoped Topology - -theorem ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ [SigmaFinite μ] - {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) - (h : ∀ s, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ ≤ ∫⁻ x in s, g x ∂μ) : - f ≤ᵐ[μ] g := by - have A : ∀ (ε N : ℝ≥0) (p : ℕ), 0 < ε → - μ ({x | g x + ε ≤ f x ∧ g x ≤ N} ∩ spanningSets μ p) = 0 := by - intro ε N p εpos - let s := {x | g x + ε ≤ f x ∧ g x ≤ N} ∩ spanningSets μ p - have s_lt_top : μ s < ∞ := - (measure_mono (Set.inter_subset_right)).trans_lt (measure_spanningSets_lt_top μ p) - have A : (∫⁻ x in s, g x ∂μ) + ε * μ s ≤ (∫⁻ x in s, g x ∂μ) + 0 := - calc - (∫⁻ x in s, g x ∂μ) + ε * μ s = (∫⁻ x in s, g x ∂μ) + ∫⁻ _ in s, ε ∂μ := by - simp only [lintegral_const, Set.univ_inter, MeasurableSet.univ, Measure.restrict_apply] - _ = ∫⁻ x in s, g x + ε ∂μ := (lintegral_add_right _ measurable_const).symm - _ ≤ ∫⁻ x in s, f x ∂μ := - setLIntegral_mono_ae hf.restrict <| ae_of_all _ fun x hx => hx.1.1 - _ ≤ (∫⁻ x in s, g x ∂μ) + 0 := by - rw [add_zero, ← Measure.restrict_toMeasurable s_lt_top.ne] - refine h _ (measurableSet_toMeasurable ..) ?_ - rwa [measure_toMeasurable] - have B : (∫⁻ x in s, g x ∂μ) ≠ ∞ := - (setLIntegral_lt_top_of_le_nnreal s_lt_top.ne ⟨N, fun _ h ↦ h.1.2⟩).ne - have : (ε : ℝ≥0∞) * μ s ≤ 0 := ENNReal.le_of_add_le_add_left B A - simpa only [ENNReal.coe_eq_zero, nonpos_iff_eq_zero, mul_eq_zero, εpos.ne', false_or] - obtain ⟨u, _, u_pos, u_lim⟩ : - ∃ u : ℕ → ℝ≥0, StrictAnti u ∧ (∀ n, 0 < u n) ∧ Tendsto u atTop (𝓝 0) := - exists_seq_strictAnti_tendsto (0 : ℝ≥0) - let s := fun n : ℕ => {x | g x + u n ≤ f x ∧ g x ≤ (n : ℝ≥0)} ∩ spanningSets μ n - have μs : ∀ n, μ (s n) = 0 := fun n => A _ _ _ (u_pos n) - have B : {x | f x ≤ g x}ᶜ ⊆ ⋃ n, s n := by - intro x hx - simp only [Set.mem_compl_iff, Set.mem_setOf, not_le] at hx - have L1 : ∀ᶠ n in atTop, g x + u n ≤ f x := by - have : Tendsto (fun n => g x + u n) atTop (𝓝 (g x + (0 : ℝ≥0))) := - tendsto_const_nhds.add (ENNReal.tendsto_coe.2 u_lim) - simp only [ENNReal.coe_zero, add_zero] at this - exact this.eventually_le_const hx - have L2 : ∀ᶠ n : ℕ in (atTop : Filter ℕ), g x ≤ (n : ℝ≥0) := - have : Tendsto (fun n : ℕ => ((n : ℝ≥0) : ℝ≥0∞)) atTop (𝓝 ∞) := by - simp only [ENNReal.coe_natCast] - exact ENNReal.tendsto_nat_nhds_top - this.eventually_const_le (hx.trans_le le_top) - apply Set.mem_iUnion.2 - exact ((L1.and L2).and (eventually_mem_spanningSets μ x)).exists - refine le_antisymm ?_ bot_le - calc - μ {x : α | (fun x : α => f x ≤ g x) x}ᶜ ≤ μ (⋃ n, s n) := measure_mono B - _ ≤ ∑' n, μ (s n) := measure_iUnion_le _ - _ = 0 := by simp only [μs, tsum_zero] - -@[deprecated (since := "2024-06-29")] -alias ae_le_of_forall_set_lintegral_le_of_sigmaFinite₀ := - ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ - -theorem ae_le_of_forall_setLIntegral_le_of_sigmaFinite [SigmaFinite μ] {f g : α → ℝ≥0∞} - (hf : Measurable f) - (h : ∀ s, MeasurableSet s → μ s < ∞ → (∫⁻ x in s, f x ∂μ) ≤ ∫⁻ x in s, g x ∂μ) : f ≤ᵐ[μ] g := - ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ hf.aemeasurable h - -@[deprecated (since := "2024-06-29")] -alias ae_le_of_forall_set_lintegral_le_of_sigmaFinite := - ae_le_of_forall_setLIntegral_le_of_sigmaFinite - -theorem ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite₀ [SigmaFinite μ] - {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) (hg : AEMeasurable g μ) - (h : ∀ s, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ) : f =ᵐ[μ] g := by - have A : f ≤ᵐ[μ] g := - ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ hf fun s hs h's => le_of_eq (h s hs h's) - have B : g ≤ᵐ[μ] f := - ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ hg fun s hs h's => ge_of_eq (h s hs h's) - filter_upwards [A, B] with x using le_antisymm - -@[deprecated (since := "2024-06-29")] -alias ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite₀ := - ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite₀ - -theorem ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite [SigmaFinite μ] {f g : α → ℝ≥0∞} - (hf : Measurable f) (hg : Measurable g) - (h : ∀ s, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ) : f =ᵐ[μ] g := - ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite₀ hf.aemeasurable hg.aemeasurable h - -@[deprecated (since := "2024-06-29")] -alias ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite := - ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite - -end ENNReal - section Real variable {f : α → ℝ} @@ -642,55 +516,4 @@ lemma ae_eq_zero_of_forall_setIntegral_isCompact_eq_zero' end AeEqOfForallSetIntegralEq -section Lintegral - -theorem AEMeasurable.ae_eq_of_forall_setLIntegral_eq {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) - (hg : AEMeasurable g μ) (hfi : ∫⁻ x, f x ∂μ ≠ ∞) (hgi : ∫⁻ x, g x ∂μ ≠ ∞) - (hfg : ∀ ⦃s⦄, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ) : - f =ᵐ[μ] g := by - refine - ENNReal.eventuallyEq_of_toReal_eventuallyEq (ae_lt_top' hf hfi).ne_of_lt - (ae_lt_top' hg hgi).ne_of_lt - (Integrable.ae_eq_of_forall_setIntegral_eq _ _ - (integrable_toReal_of_lintegral_ne_top hf hfi) - (integrable_toReal_of_lintegral_ne_top hg hgi) fun s hs hs' => ?_) - rw [integral_eq_lintegral_of_nonneg_ae, integral_eq_lintegral_of_nonneg_ae] - · congr 1 - rw [lintegral_congr_ae (ofReal_toReal_ae_eq _), lintegral_congr_ae (ofReal_toReal_ae_eq _)] - · exact hfg hs hs' - · refine ae_lt_top' hg.restrict (ne_of_lt (lt_of_le_of_lt ?_ hgi.lt_top)) - exact @setLIntegral_univ α _ μ g ▸ lintegral_mono_set (Set.subset_univ _) - · refine ae_lt_top' hf.restrict (ne_of_lt (lt_of_le_of_lt ?_ hfi.lt_top)) - exact @setLIntegral_univ α _ μ f ▸ lintegral_mono_set (Set.subset_univ _) - -- putting the proofs where they are used is extremely slow - exacts [ae_of_all _ fun x => ENNReal.toReal_nonneg, - hg.ennreal_toReal.restrict.aestronglyMeasurable, ae_of_all _ fun x => ENNReal.toReal_nonneg, - hf.ennreal_toReal.restrict.aestronglyMeasurable] - -@[deprecated (since := "2024-06-29")] -alias AEMeasurable.ae_eq_of_forall_set_lintegral_eq := AEMeasurable.ae_eq_of_forall_setLIntegral_eq - -end Lintegral - -section WithDensity - -variable {m : MeasurableSpace α} {μ : Measure α} - -theorem withDensity_eq_iff_of_sigmaFinite [SigmaFinite μ] {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) - (hg : AEMeasurable g μ) : μ.withDensity f = μ.withDensity g ↔ f =ᵐ[μ] g := - ⟨fun hfg ↦ by - refine ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite₀ hf hg fun s hs _ ↦ ?_ - rw [← withDensity_apply f hs, ← withDensity_apply g hs, ← hfg], withDensity_congr_ae⟩ - -theorem withDensity_eq_iff {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) - (hg : AEMeasurable g μ) (hfi : ∫⁻ x, f x ∂μ ≠ ∞) : - μ.withDensity f = μ.withDensity g ↔ f =ᵐ[μ] g := - ⟨fun hfg ↦ by - refine AEMeasurable.ae_eq_of_forall_setLIntegral_eq hf hg hfi ?_ fun s hs _ ↦ ?_ - · rwa [← setLIntegral_univ, ← withDensity_apply g MeasurableSet.univ, ← hfg, - withDensity_apply f MeasurableSet.univ, setLIntegral_univ] - · rw [← withDensity_apply f hs, ← withDensity_apply g hs, ← hfg], withDensity_congr_ae⟩ - -end WithDensity - end MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/AEEqOfLIntegral.lean b/Mathlib/MeasureTheory/Function/AEEqOfLIntegral.lean new file mode 100644 index 0000000000000..6e2b5cedfb464 --- /dev/null +++ b/Mathlib/MeasureTheory/Function/AEEqOfLIntegral.lean @@ -0,0 +1,210 @@ +/- +Copyright (c) 2021 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.Function.StronglyMeasurable.ENNReal +import Mathlib.MeasureTheory.Measure.WithDensity + +/-! # From equality of integrals to equality of functions + +This file provides various statements of the general form "if two functions have the same integral +on all sets, then they are equal almost everywhere". +The different lemmas use various hypotheses on the class of functions, on the target space or on the +possible finiteness of the measure. + +This file is about Lebesgue integrals. See the file `AEEqOfIntegral` for Bochner integrals. + +## Main statements + +The results listed below apply to two functions `f, g`, under the hypothesis that +for all measurable sets `s` with finite measure, `∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ`. +The conclusion is then `f =ᵐ[μ] g`. The main lemmas are: +* `ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite`: case of a sigma-finite measure. +* `AEMeasurable.ae_eq_of_forall_setLIntegral_eq`: for functions which are `AEMeasurable` and + have finite integral. + +-/ + + +open Filter + +open scoped ENNReal NNReal MeasureTheory Topology + +namespace MeasureTheory + +variable {α : Type*} {m m0 : MeasurableSpace α} {μ : Measure α} {p : ℝ≥0∞} + +theorem ae_const_le_iff_forall_lt_measure_zero {β} [LinearOrder β] [TopologicalSpace β] + [OrderTopology β] [FirstCountableTopology β] (f : α → β) (c : β) : + (∀ᵐ x ∂μ, c ≤ f x) ↔ ∀ b < c, μ {x | f x ≤ b} = 0 := by + rw [ae_iff] + push_neg + constructor + · intro h b hb + exact measure_mono_null (fun y hy => (lt_of_le_of_lt hy hb : _)) h + intro hc + by_cases h : ∀ b, c ≤ b + · have : {a : α | f a < c} = ∅ := by + apply Set.eq_empty_iff_forall_not_mem.2 fun x hx => ?_ + exact (lt_irrefl _ (lt_of_lt_of_le hx (h (f x)))).elim + simp [this] + by_cases H : ¬IsLUB (Set.Iio c) c + · have : c ∈ upperBounds (Set.Iio c) := fun y hy => le_of_lt hy + obtain ⟨b, b_up, bc⟩ : ∃ b : β, b ∈ upperBounds (Set.Iio c) ∧ b < c := by + simpa [IsLUB, IsLeast, this, lowerBounds] using H + exact measure_mono_null (fun x hx => b_up hx) (hc b bc) + push_neg at H h + obtain ⟨u, _, u_lt, u_lim, -⟩ : + ∃ u : ℕ → β, + StrictMono u ∧ (∀ n : ℕ, u n < c) ∧ Tendsto u atTop (𝓝 c) ∧ ∀ n : ℕ, u n ∈ Set.Iio c := + H.exists_seq_strictMono_tendsto_of_not_mem (lt_irrefl c) h + have h_Union : {x | f x < c} = ⋃ n : ℕ, {x | f x ≤ u n} := by + ext1 x + simp_rw [Set.mem_iUnion, Set.mem_setOf_eq] + constructor <;> intro h + · obtain ⟨n, hn⟩ := ((tendsto_order.1 u_lim).1 _ h).exists; exact ⟨n, hn.le⟩ + · obtain ⟨n, hn⟩ := h; exact hn.trans_lt (u_lt _) + rw [h_Union, measure_iUnion_null_iff] + intro n + exact hc _ (u_lt n) + +theorem ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ [SigmaFinite μ] + {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) + (h : ∀ s, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ ≤ ∫⁻ x in s, g x ∂μ) : + f ≤ᵐ[μ] g := by + have A : ∀ (ε N : ℝ≥0) (p : ℕ), 0 < ε → + μ ({x | g x + ε ≤ f x ∧ g x ≤ N} ∩ spanningSets μ p) = 0 := by + intro ε N p εpos + let s := {x | g x + ε ≤ f x ∧ g x ≤ N} ∩ spanningSets μ p + have s_lt_top : μ s < ∞ := + (measure_mono (Set.inter_subset_right)).trans_lt (measure_spanningSets_lt_top μ p) + have A : (∫⁻ x in s, g x ∂μ) + ε * μ s ≤ (∫⁻ x in s, g x ∂μ) + 0 := + calc + (∫⁻ x in s, g x ∂μ) + ε * μ s = (∫⁻ x in s, g x ∂μ) + ∫⁻ _ in s, ε ∂μ := by + simp only [lintegral_const, Set.univ_inter, MeasurableSet.univ, Measure.restrict_apply] + _ = ∫⁻ x in s, g x + ε ∂μ := (lintegral_add_right _ measurable_const).symm + _ ≤ ∫⁻ x in s, f x ∂μ := + setLIntegral_mono_ae hf.restrict <| ae_of_all _ fun x hx => hx.1.1 + _ ≤ (∫⁻ x in s, g x ∂μ) + 0 := by + rw [add_zero, ← Measure.restrict_toMeasurable s_lt_top.ne] + refine h _ (measurableSet_toMeasurable ..) ?_ + rwa [measure_toMeasurable] + have B : (∫⁻ x in s, g x ∂μ) ≠ ∞ := + (setLIntegral_lt_top_of_le_nnreal s_lt_top.ne ⟨N, fun _ h ↦ h.1.2⟩).ne + have : (ε : ℝ≥0∞) * μ s ≤ 0 := ENNReal.le_of_add_le_add_left B A + simpa only [ENNReal.coe_eq_zero, nonpos_iff_eq_zero, mul_eq_zero, εpos.ne', false_or] + obtain ⟨u, _, u_pos, u_lim⟩ : + ∃ u : ℕ → ℝ≥0, StrictAnti u ∧ (∀ n, 0 < u n) ∧ Tendsto u atTop (𝓝 0) := + exists_seq_strictAnti_tendsto (0 : ℝ≥0) + let s := fun n : ℕ => {x | g x + u n ≤ f x ∧ g x ≤ (n : ℝ≥0)} ∩ spanningSets μ n + have μs : ∀ n, μ (s n) = 0 := fun n => A _ _ _ (u_pos n) + have B : {x | f x ≤ g x}ᶜ ⊆ ⋃ n, s n := by + intro x hx + simp only [Set.mem_compl_iff, Set.mem_setOf, not_le] at hx + have L1 : ∀ᶠ n in atTop, g x + u n ≤ f x := by + have : Tendsto (fun n => g x + u n) atTop (𝓝 (g x + (0 : ℝ≥0))) := + tendsto_const_nhds.add (ENNReal.tendsto_coe.2 u_lim) + simp only [ENNReal.coe_zero, add_zero] at this + exact this.eventually_le_const hx + have L2 : ∀ᶠ n : ℕ in (atTop : Filter ℕ), g x ≤ (n : ℝ≥0) := + have : Tendsto (fun n : ℕ => ((n : ℝ≥0) : ℝ≥0∞)) atTop (𝓝 ∞) := by + simp only [ENNReal.coe_natCast] + exact ENNReal.tendsto_nat_nhds_top + this.eventually_const_le (hx.trans_le le_top) + apply Set.mem_iUnion.2 + exact ((L1.and L2).and (eventually_mem_spanningSets μ x)).exists + refine le_antisymm ?_ bot_le + calc + μ {x : α | (fun x : α => f x ≤ g x) x}ᶜ ≤ μ (⋃ n, s n) := measure_mono B + _ ≤ ∑' n, μ (s n) := measure_iUnion_le _ + _ = 0 := by simp only [μs, tsum_zero] + +@[deprecated (since := "2024-06-29")] +alias ae_le_of_forall_set_lintegral_le_of_sigmaFinite₀ := + ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ + +theorem ae_le_of_forall_setLIntegral_le_of_sigmaFinite [SigmaFinite μ] {f g : α → ℝ≥0∞} + (hf : Measurable f) + (h : ∀ s, MeasurableSet s → μ s < ∞ → (∫⁻ x in s, f x ∂μ) ≤ ∫⁻ x in s, g x ∂μ) : f ≤ᵐ[μ] g := + ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ hf.aemeasurable h + +@[deprecated (since := "2024-06-29")] +alias ae_le_of_forall_set_lintegral_le_of_sigmaFinite := + ae_le_of_forall_setLIntegral_le_of_sigmaFinite + +theorem ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite₀ [SigmaFinite μ] + {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) (hg : AEMeasurable g μ) + (h : ∀ s, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ) : f =ᵐ[μ] g := by + have A : f ≤ᵐ[μ] g := + ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ hf fun s hs h's => le_of_eq (h s hs h's) + have B : g ≤ᵐ[μ] f := + ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ hg fun s hs h's => ge_of_eq (h s hs h's) + filter_upwards [A, B] with x using le_antisymm + +@[deprecated (since := "2024-06-29")] +alias ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite₀ := + ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite₀ + +theorem ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite [SigmaFinite μ] {f g : α → ℝ≥0∞} + (hf : Measurable f) (hg : Measurable g) + (h : ∀ s, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ) : f =ᵐ[μ] g := + ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite₀ hf.aemeasurable hg.aemeasurable h + +@[deprecated (since := "2024-06-29")] +alias ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite := + ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite + +theorem AEMeasurable.ae_eq_of_forall_setLIntegral_eq {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) + (hg : AEMeasurable g μ) (hfi : ∫⁻ x, f x ∂μ ≠ ∞) (hgi : ∫⁻ x, g x ∂μ ≠ ∞) + (hfg : ∀ ⦃s⦄, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ) : + f =ᵐ[μ] g := by + have hf' : AEFinStronglyMeasurable f μ := + ENNReal.aefinStronglyMeasurable_of_aemeasurable hfi hf + have hg' : AEFinStronglyMeasurable g μ := + ENNReal.aefinStronglyMeasurable_of_aemeasurable hgi hg + let s := hf'.sigmaFiniteSet + let t := hg'.sigmaFiniteSet + suffices f =ᵐ[μ.restrict (s ∪ t)] g by + refine ae_of_ae_restrict_of_ae_restrict_compl _ this ?_ + simp only [Set.compl_union] + have h1 : f =ᵐ[μ.restrict sᶜ] 0 := hf'.ae_eq_zero_compl + have h2 : g =ᵐ[μ.restrict tᶜ] 0 := hg'.ae_eq_zero_compl + rw [ae_restrict_iff' (hf'.measurableSet.compl.inter hg'.measurableSet.compl)] + rw [EventuallyEq, ae_restrict_iff' hf'.measurableSet.compl] at h1 + rw [EventuallyEq, ae_restrict_iff' hg'.measurableSet.compl] at h2 + filter_upwards [h1, h2] with x h1 h2 hx + rw [h1 (Set.inter_subset_left hx), h2 (Set.inter_subset_right hx)] + have := hf'.sigmaFinite_restrict + have := hg'.sigmaFinite_restrict + refine ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite₀ hf.restrict hg.restrict + fun u hu huμ ↦ ?_ + rw [Measure.restrict_restrict hu] + rw [Measure.restrict_apply hu] at huμ + exact hfg (hu.inter (hf'.measurableSet.union hg'.measurableSet)) huμ + +@[deprecated (since := "2024-06-29")] +alias AEMeasurable.ae_eq_of_forall_set_lintegral_eq := AEMeasurable.ae_eq_of_forall_setLIntegral_eq + +section WithDensity + +variable {m : MeasurableSpace α} {μ : Measure α} + +theorem withDensity_eq_iff_of_sigmaFinite [SigmaFinite μ] {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) + (hg : AEMeasurable g μ) : μ.withDensity f = μ.withDensity g ↔ f =ᵐ[μ] g := + ⟨fun hfg ↦ by + refine ae_eq_of_forall_setLIntegral_eq_of_sigmaFinite₀ hf hg fun s hs _ ↦ ?_ + rw [← withDensity_apply f hs, ← withDensity_apply g hs, ← hfg], withDensity_congr_ae⟩ + +theorem withDensity_eq_iff {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) + (hg : AEMeasurable g μ) (hfi : ∫⁻ x, f x ∂μ ≠ ∞) : + μ.withDensity f = μ.withDensity g ↔ f =ᵐ[μ] g := + ⟨fun hfg ↦ by + refine AEMeasurable.ae_eq_of_forall_setLIntegral_eq hf hg hfi ?_ fun s hs _ ↦ ?_ + · rwa [← setLIntegral_univ, ← withDensity_apply g MeasurableSet.univ, ← hfg, + withDensity_apply f MeasurableSet.univ, setLIntegral_univ] + · rw [← withDensity_apply f hs, ← withDensity_apply g hs, ← hfg], withDensity_congr_ae⟩ + +end WithDensity + +end MeasureTheory From 8d47a00443ecfd3ed7c1e5540730cd77d7a27956 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Miyahara=20K=C5=8D?= Date: Sat, 4 Jan 2025 13:18:11 +0000 Subject: [PATCH 136/189] chore: `inherit_doc`s for notations (#20376) --- Mathlib/Algebra/Symmetrized.lean | 2 +- Mathlib/Analysis/Normed/Affine/Isometry.lean | 2 +- Mathlib/Data/Real/Cardinality.lean | 2 +- Mathlib/Data/Seq/WSeq.lean | 2 +- Mathlib/Data/Stream/Defs.lean | 8 +++--- Mathlib/Geometry/Manifold/Algebra/Monoid.lean | 4 +-- Mathlib/LinearAlgebra/CrossProduct.lean | 2 +- Mathlib/LinearAlgebra/SModEq.lean | 2 +- Mathlib/Logic/Function/Defs.lean | 2 +- .../Integral/IntervalIntegral.lean | 4 +++ Mathlib/ModelTheory/LanguageMap.lean | 2 +- Mathlib/ModelTheory/Syntax.lean | 8 +++--- Mathlib/NumberTheory/Zsqrtd/Basic.lean | 2 +- Mathlib/SetTheory/Cardinal/Continuum.lean | 2 +- .../Topology/ContinuousMap/Bounded/Basic.lean | 2 +- Mathlib/Topology/Hom/Open.lean | 2 +- Mathlib/Topology/Homotopy/Equiv.lean | 2 +- Mathlib/Topology/MetricSpace/Dilation.lean | 2 +- .../Topology/MetricSpace/DilationEquiv.lean | 2 +- Mathlib/Topology/Order/Hom/Basic.lean | 2 +- scripts/nolints.json | 27 ------------------- 21 files changed, 30 insertions(+), 53 deletions(-) diff --git a/Mathlib/Algebra/Symmetrized.lean b/Mathlib/Algebra/Symmetrized.lean index 45a69d0768566..5d4e197bf27e2 100644 --- a/Mathlib/Algebra/Symmetrized.lean +++ b/Mathlib/Algebra/Symmetrized.lean @@ -35,7 +35,7 @@ open Function def SymAlg (α : Type*) : Type _ := α -postfix:max "ˢʸᵐ" => SymAlg +@[inherit_doc] postfix:max "ˢʸᵐ" => SymAlg namespace SymAlg diff --git a/Mathlib/Analysis/Normed/Affine/Isometry.lean b/Mathlib/Analysis/Normed/Affine/Isometry.lean index ad879f1396054..e3fbc948e36bf 100644 --- a/Mathlib/Analysis/Normed/Affine/Isometry.lean +++ b/Mathlib/Analysis/Normed/Affine/Isometry.lean @@ -267,7 +267,7 @@ structure AffineIsometryEquiv extends P ≃ᵃ[𝕜] P₂ where variable {𝕜 P P₂} -- `≃ᵃᵢ` would be more consistent with the linear isometry equiv notation, but it is uglier -notation:25 P " ≃ᵃⁱ[" 𝕜:25 "] " P₂:0 => AffineIsometryEquiv 𝕜 P P₂ +@[inherit_doc] notation:25 P " ≃ᵃⁱ[" 𝕜:25 "] " P₂:0 => AffineIsometryEquiv 𝕜 P P₂ namespace AffineIsometryEquiv diff --git a/Mathlib/Data/Real/Cardinality.lean b/Mathlib/Data/Real/Cardinality.lean index a4bda74e874be..bba754011be99 100644 --- a/Mathlib/Data/Real/Cardinality.lean +++ b/Mathlib/Data/Real/Cardinality.lean @@ -34,7 +34,7 @@ We conclude that all intervals with distinct endpoints have cardinality continuu ## Notation -* `𝔠` : notation for `Cardinal.Continuum` in locale `Cardinal`, defined in `SetTheory.Continuum`. +* `𝔠` : notation for `Cardinal.continuum` in locale `Cardinal`, defined in `SetTheory.Continuum`. ## Tags continuum, cardinality, reals, cardinality of the reals diff --git a/Mathlib/Data/Seq/WSeq.lean b/Mathlib/Data/Seq/WSeq.lean index 5b57888ef3f6e..e860eac5943d9 100644 --- a/Mathlib/Data/Seq/WSeq.lean +++ b/Mathlib/Data/Seq/WSeq.lean @@ -454,7 +454,7 @@ theorem liftRel_destruct_iff {R : α → β → Prop} {s : WSeq α} {t : WSeq β apply Or.inl⟩⟩ -- Porting note: To avoid ambiguous notation, `~` became `~ʷ`. -infixl:50 " ~ʷ " => Equiv +@[inherit_doc] infixl:50 " ~ʷ " => Equiv theorem destruct_congr {s t : WSeq α} : s ~ʷ t → Computation.LiftRel (BisimO (· ~ʷ ·)) (destruct s) (destruct t) := diff --git a/Mathlib/Data/Stream/Defs.lean b/Mathlib/Data/Stream/Defs.lean index 6a1171545ddf4..5d5ca8300b137 100644 --- a/Mathlib/Data/Stream/Defs.lean +++ b/Mathlib/Data/Stream/Defs.lean @@ -26,7 +26,7 @@ def cons (a : α) (s : Stream' α) : Stream' α | 0 => a | n + 1 => s n -scoped infixr:67 " :: " => cons +@[inherit_doc] scoped infixr:67 " :: " => cons /-- Get the `n`-th element of a stream. -/ def get (s : Stream' α) (n : ℕ) : α := s n @@ -90,7 +90,7 @@ abbrev unfolds (g : α → β) (f : α → α) (a : α) : Stream' β := def interleave (s₁ s₂ : Stream' α) : Stream' α := corecOn (s₁, s₂) (fun ⟨s₁, _⟩ => head s₁) fun ⟨s₁, s₂⟩ => (s₂, tail s₁) -infixl:65 " ⋈ " => interleave +@[inherit_doc] infixl:65 " ⋈ " => interleave /-- Elements of a stream with even indices. -/ def even (s : Stream' α) : Stream' α := @@ -105,7 +105,7 @@ def appendStream' : List α → Stream' α → Stream' α | [], s => s | List.cons a l, s => a::appendStream' l s -infixl:65 " ++ₛ " => appendStream' +@[inherit_doc] infixl:65 " ++ₛ " => appendStream' /-- `take n s` returns a list of the `n` first elements of stream `s` -/ def take : ℕ → Stream' α → List α @@ -147,7 +147,7 @@ def pure (a : α) : Stream' α := /-- Given a stream of functions and a stream of values, apply `n`-th function to `n`-th value. -/ def apply (f : Stream' (α → β)) (s : Stream' α) : Stream' β := fun n => (get f n) (get s n) -infixl:75 " ⊛ " => apply +@[inherit_doc] infixl:75 " ⊛ " => apply -- Porting note: "input as \o*" was here but doesn't work for the above notation /-- The stream of natural numbers: `Stream'.get n Stream'.nats = n`. -/ diff --git a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean index 7e6a90b613a22..dba923efb9653 100644 --- a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean +++ b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean @@ -177,10 +177,10 @@ def smoothRightMul : C^∞⟮I, G; I, G⟯ := ⟨rightMul g, contMDiff_mul_right⟩ -- Left multiplication. The abbreviation is `MIL`. -scoped[LieGroup] notation "𝑳" => smoothLeftMul +@[inherit_doc] scoped[LieGroup] notation "𝑳" => smoothLeftMul -- Right multiplication. The abbreviation is `MIR`. -scoped[LieGroup] notation "𝑹" => smoothRightMul +@[inherit_doc] scoped[LieGroup] notation "𝑹" => smoothRightMul open scoped LieGroup diff --git a/Mathlib/LinearAlgebra/CrossProduct.lean b/Mathlib/LinearAlgebra/CrossProduct.lean index 913de0b9539bc..b1b889945b83e 100644 --- a/Mathlib/LinearAlgebra/CrossProduct.lean +++ b/Mathlib/LinearAlgebra/CrossProduct.lean @@ -55,7 +55,7 @@ def crossProduct : (Fin 3 → R) →ₗ[R] (Fin 3 → R) →ₗ[R] Fin 3 → R : · intros simp_rw [smul_vec3, Pi.smul_apply, smul_sub, mul_smul_comm] -scoped[Matrix] infixl:74 " ×₃ " => crossProduct +@[inherit_doc] scoped[Matrix] infixl:74 " ×₃ " => crossProduct theorem cross_apply (a b : Fin 3 → R) : a ×₃ b = ![a 1 * b 2 - a 2 * b 1, a 2 * b 0 - a 0 * b 2, a 0 * b 1 - a 1 * b 0] := rfl diff --git a/Mathlib/LinearAlgebra/SModEq.lean b/Mathlib/LinearAlgebra/SModEq.lean index fe70f13aff7a2..9cda07bf6a5c2 100644 --- a/Mathlib/LinearAlgebra/SModEq.lean +++ b/Mathlib/LinearAlgebra/SModEq.lean @@ -26,7 +26,7 @@ variable {N : Type*} [AddCommGroup N] [Module R N] (V V₁ V₂ : Submodule R N) def SModEq (x y : M) : Prop := (Submodule.Quotient.mk x : M ⧸ U) = Submodule.Quotient.mk y -notation:50 x " ≡ " y " [SMOD " N "]" => SModEq N x y +@[inherit_doc] notation:50 x " ≡ " y " [SMOD " N "]" => SModEq N x y variable {U U₁ U₂} diff --git a/Mathlib/Logic/Function/Defs.lean b/Mathlib/Logic/Function/Defs.lean index 6d82f855d2f0c..21054aab26345 100644 --- a/Mathlib/Logic/Function/Defs.lean +++ b/Mathlib/Logic/Function/Defs.lean @@ -36,7 +36,7 @@ and type of `f (g x)` depends on `x` and `g x`. -/ def dcomp {β : α → Sort u₂} {φ : ∀ {x : α}, β x → Sort u₃} (f : ∀ {x : α} (y : β x), φ y) (g : ∀ x, β x) : ∀ x, φ (g x) := fun x => f (g x) -infixr:80 " ∘' " => Function.dcomp +@[inherit_doc] infixr:80 " ∘' " => Function.dcomp /-- Given functions `f : β → β → φ` and `g : α → β`, produce a function `α → α → φ` that evaluates `g` on each argument, then applies `f` to the results. Can be used, e.g., to transfer a relation diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index 67c398ceac956..2e2126e29a549 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -420,8 +420,12 @@ as `∫ x in Ioc a b, f x ∂μ - ∫ x in Ioc b a, f x ∂μ`. If `a ≤ b`, th def intervalIntegral (f : ℝ → E) (a b : ℝ) (μ : Measure ℝ) : E := (∫ x in Ioc a b, f x ∂μ) - ∫ x in Ioc b a, f x ∂μ +@[inherit_doc intervalIntegral] notation3"∫ "(...)" in "a".."b", "r:60:(scoped f => f)" ∂"μ:70 => intervalIntegral r a b μ +/-- 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 +`∫ x in Ioc a b, f x`, otherwise it equals `-∫ x in Ioc b a, f x`. -/ notation3"∫ "(...)" in "a".."b", "r:60:(scoped f => intervalIntegral f a b volume) => r namespace intervalIntegral diff --git a/Mathlib/ModelTheory/LanguageMap.lean b/Mathlib/ModelTheory/LanguageMap.lean index e7a57200c37cb..2daf318f51f3e 100644 --- a/Mathlib/ModelTheory/LanguageMap.lean +++ b/Mathlib/ModelTheory/LanguageMap.lean @@ -278,7 +278,7 @@ structure LEquiv (L L' : Language) where left_inv : invLHom.comp toLHom = LHom.id L right_inv : toLHom.comp invLHom = LHom.id L' -infixl:10 " ≃ᴸ " => LEquiv +@[inherit_doc] infixl:10 " ≃ᴸ " => LEquiv -- \^L namespace LEquiv diff --git a/Mathlib/ModelTheory/Syntax.lean b/Mathlib/ModelTheory/Syntax.lean index cff9e39352aa6..ca64daf425250 100644 --- a/Mathlib/ModelTheory/Syntax.lean +++ b/Mathlib/ModelTheory/Syntax.lean @@ -709,7 +709,7 @@ def onSentence (φ : L ≃ᴸ L') : L.Sentence ≃ L'.Sentence := end LEquiv -scoped[FirstOrder] infixl:88 " =' " => FirstOrder.Language.Term.bdEqual +@[inherit_doc] scoped[FirstOrder] infixl:88 " =' " => FirstOrder.Language.Term.bdEqual -- input \~- or \simeq scoped[FirstOrder] infixr:62 " ⟹ " => FirstOrder.Language.BoundedFormula.imp @@ -717,13 +717,13 @@ scoped[FirstOrder] infixr:62 " ⟹ " => FirstOrder.Language.BoundedFormula.imp scoped[FirstOrder] prefix:110 "∀'" => FirstOrder.Language.BoundedFormula.all -scoped[FirstOrder] prefix:arg "∼" => FirstOrder.Language.BoundedFormula.not +@[inherit_doc] scoped[FirstOrder] prefix:arg "∼" => FirstOrder.Language.BoundedFormula.not -- input \~, the ASCII character ~ has too low precedence -scoped[FirstOrder] infixl:61 " ⇔ " => FirstOrder.Language.BoundedFormula.iff +@[inherit_doc] scoped[FirstOrder] infixl:61 " ⇔ " => FirstOrder.Language.BoundedFormula.iff -- input \<=> -scoped[FirstOrder] prefix:110 "∃'" => FirstOrder.Language.BoundedFormula.ex +@[inherit_doc] scoped[FirstOrder] prefix:110 "∃'" => FirstOrder.Language.BoundedFormula.ex -- input \ex namespace Formula diff --git a/Mathlib/NumberTheory/Zsqrtd/Basic.lean b/Mathlib/NumberTheory/Zsqrtd/Basic.lean index b7c9c0db4e4a0..9d477daf4ffca 100644 --- a/Mathlib/NumberTheory/Zsqrtd/Basic.lean +++ b/Mathlib/NumberTheory/Zsqrtd/Basic.lean @@ -31,7 +31,7 @@ structure Zsqrtd (d : ℤ) where im : ℤ deriving DecidableEq -prefix:100 "ℤ√" => Zsqrtd +@[inherit_doc] prefix:100 "ℤ√" => Zsqrtd namespace Zsqrtd diff --git a/Mathlib/SetTheory/Cardinal/Continuum.lean b/Mathlib/SetTheory/Cardinal/Continuum.lean index fd2ba0a5e9087..c04a55eca2aa0 100644 --- a/Mathlib/SetTheory/Cardinal/Continuum.lean +++ b/Mathlib/SetTheory/Cardinal/Continuum.lean @@ -27,7 +27,7 @@ open Cardinal def continuum : Cardinal.{u} := 2 ^ ℵ₀ -scoped notation "𝔠" => Cardinal.continuum +@[inherit_doc] scoped notation "𝔠" => Cardinal.continuum @[simp] theorem two_power_aleph0 : 2 ^ ℵ₀ = 𝔠 := diff --git a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean index c32fd5163a6e0..c634039f9359d 100644 --- a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean @@ -42,7 +42,7 @@ structure BoundedContinuousFunction (α : Type u) (β : Type v) [TopologicalSpac [PseudoMetricSpace β] extends ContinuousMap α β : Type max u v where map_bounded' : ∃ C, ∀ x y, dist (toFun x) (toFun y) ≤ C -scoped[BoundedContinuousFunction] infixr:25 " →ᵇ " => BoundedContinuousFunction +@[inherit_doc] scoped[BoundedContinuousFunction] infixr:25 " →ᵇ " => BoundedContinuousFunction section diff --git a/Mathlib/Topology/Hom/Open.lean b/Mathlib/Topology/Hom/Open.lean index 86051435726a8..5daac366a7826 100644 --- a/Mathlib/Topology/Hom/Open.lean +++ b/Mathlib/Topology/Hom/Open.lean @@ -32,7 +32,7 @@ structure ContinuousOpenMap (α β : Type*) [TopologicalSpace α] [TopologicalSp ContinuousMap α β where map_open' : IsOpenMap toFun -infixr:25 " →CO " => ContinuousOpenMap +@[inherit_doc] infixr:25 " →CO " => ContinuousOpenMap section diff --git a/Mathlib/Topology/Homotopy/Equiv.lean b/Mathlib/Topology/Homotopy/Equiv.lean index 2557422e9d194..d909aa0b84ea4 100644 --- a/Mathlib/Topology/Homotopy/Equiv.lean +++ b/Mathlib/Topology/Homotopy/Equiv.lean @@ -42,7 +42,7 @@ structure HomotopyEquiv (X : Type u) (Y : Type v) [TopologicalSpace X] [Topologi left_inv : (invFun.comp toFun).Homotopic (ContinuousMap.id X) right_inv : (toFun.comp invFun).Homotopic (ContinuousMap.id Y) -scoped infixl:25 " ≃ₕ " => ContinuousMap.HomotopyEquiv +@[inherit_doc] scoped infixl:25 " ≃ₕ " => ContinuousMap.HomotopyEquiv namespace HomotopyEquiv diff --git a/Mathlib/Topology/MetricSpace/Dilation.lean b/Mathlib/Topology/MetricSpace/Dilation.lean index 6ab93147d696b..ee4cc20f7a1f3 100644 --- a/Mathlib/Topology/MetricSpace/Dilation.lean +++ b/Mathlib/Topology/MetricSpace/Dilation.lean @@ -62,7 +62,7 @@ structure Dilation where toFun : α → β edist_eq' : ∃ r : ℝ≥0, r ≠ 0 ∧ ∀ x y : α, edist (toFun x) (toFun y) = r * edist x y -infixl:25 " →ᵈ " => Dilation +@[inherit_doc] infixl:25 " →ᵈ " => Dilation /-- `DilationClass F α β r` states that `F` is a type of `r`-dilations. You should extend this typeclass when you extend `Dilation`. -/ diff --git a/Mathlib/Topology/MetricSpace/DilationEquiv.lean b/Mathlib/Topology/MetricSpace/DilationEquiv.lean index 3fffba7e08305..097fe5fbc2543 100644 --- a/Mathlib/Topology/MetricSpace/DilationEquiv.lean +++ b/Mathlib/Topology/MetricSpace/DilationEquiv.lean @@ -42,7 +42,7 @@ end Class structure DilationEquiv (X Y : Type*) [PseudoEMetricSpace X] [PseudoEMetricSpace Y] extends X ≃ Y, Dilation X Y -infixl:25 " ≃ᵈ " => DilationEquiv +@[inherit_doc] infixl:25 " ≃ᵈ " => DilationEquiv namespace DilationEquiv diff --git a/Mathlib/Topology/Order/Hom/Basic.lean b/Mathlib/Topology/Order/Hom/Basic.lean index b0e27200893c5..2f4067dad99ae 100644 --- a/Mathlib/Topology/Order/Hom/Basic.lean +++ b/Mathlib/Topology/Order/Hom/Basic.lean @@ -36,7 +36,7 @@ structure ContinuousOrderHom (α β : Type*) [Preorder α] [Preorder β] [Topolo [TopologicalSpace β] extends OrderHom α β where continuous_toFun : Continuous toFun -infixr:25 " →Co " => ContinuousOrderHom +@[inherit_doc] infixr:25 " →Co " => ContinuousOrderHom section diff --git a/scripts/nolints.json b/scripts/nolints.json index 0acd19eb7d1e5..9c3201d0d844a 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -35,16 +35,6 @@ ["docBlame", "measurability!?"], ["docBlame", "tacticCancel_denoms_"], ["docBlame", "tacticUse_finite_instance"], - ["docBlame", "«term_ˢʸᵐ»"], - ["docBlame", "«term_→CO_»"], - ["docBlame", "«term_→Co_»"], - ["docBlame", "«term_→ᵈ_»"], - ["docBlame", "«term_≃ᵃⁱ[_]_»"], - ["docBlame", "«term_≃ᵈ_»"], - ["docBlame", "«term_≡_[SMOD_]»"], - ["docBlame", "«termℤ√_»"], - ["docBlame", "«term∫_In_.._,_»"], - ["docBlame", "«term∫_In_.._,_∂_»"], ["docBlame", "«term∮_InC(_,_),_»"], ["docBlame", "«term⨍_In_.._,_»"], ["docBlame", "when"], @@ -66,7 +56,6 @@ ["docBlame", "Bimod.actLeft"], ["docBlame", "Bimod.actRight"], ["docBlame", "Bitraversable.bitraverse"], - ["docBlame", "BoundedContinuousFunction.«term_→ᵇ_»"], ["docBlame", "BoundedRandom.randomR"], ["docBlame", "ByteSlice.arr"], ["docBlame", "ByteSlice.len"], @@ -77,7 +66,6 @@ ["docBlame", "CFilter.inf"], ["docBlame", "CFilter.pt"], ["docBlame", "CancelDenoms.synthesizeUsingNormNum"], - ["docBlame", "Cardinal.term𝔠"], ["docBlame", "CategoryTheory.«term_⟶[_]_»"], ["docBlame", "ChangeOfRings.«term_⊗ₜ[_,_]_»"], ["docBlame", "Combinator.I"], @@ -90,7 +78,6 @@ ["docBlame", "ContT.monadLift"], ["docBlame", "ContT.run"], ["docBlame", "ContT.withContT"], - ["docBlame", "ContinuousMap.«term_≃ₕ_»"], ["docBlame", "Ctop.f"], ["docBlame", "Ctop.inter"], ["docBlame", "Ctop.top"], @@ -141,14 +128,9 @@ ["docBlame", "FinBoolAlg.toBoolAlg"], ["docBlame", "FinPartOrd.toPartOrd"], ["docBlame", "FirstOrder.«term&_»"], - ["docBlame", "FirstOrder.«term_='_»"], - ["docBlame", "FirstOrder.«term_⇔_»"], ["docBlame", "FirstOrder.«term_⟹_»"], ["docBlame", "FirstOrder.«term∀'_»"], - ["docBlame", "FirstOrder.«term∃'_»"], - ["docBlame", "FirstOrder.«term∼_»"], ["docBlame", "Function.swap"], - ["docBlame", "Function.«term_∘'_»"], ["docBlame", "HSpace.e"], ["docBlame", "HSpace.eHmul"], ["docBlame", "HSpace.hmul"], @@ -189,8 +171,6 @@ ["docBlame", "Kronecker.«term_⊗ₖₜ[_]_»"], ["docBlame", "Kronecker.«term_⊗ₖₜ_»"], ["docBlame", "Lean.ExportM"], - ["docBlame", "LieGroup.«term𝑳»"], - ["docBlame", "LieGroup.«term𝑹»"], ["docBlame", "LinearPMap.domain"], ["docBlame", "LinearPMap.sSup"], ["docBlame", "LinearPMap.«term_†»"], @@ -203,7 +183,6 @@ ["docBlame", "Manifold.«term_≫ₕ_»"], ["docBlame", "Manifold.«term𝒅»"], ["docBlame", "Manifold.«term𝒅ₕ»"], - ["docBlame", "Matrix.«term_×₃_»"], ["docBlame", "Matrix.«term_⊕ᵥ_»"], ["docBlame", "Matrix.«term_⊙_»"], ["docBlame", "MaximalSpectrum.asIdeal"], @@ -319,10 +298,6 @@ ["docBlame", "Stream'.corec"], ["docBlame", "Stream'.corec'"], ["docBlame", "Stream'.corecOn"], - ["docBlame", "Stream'.«term_++ₛ_»"], - ["docBlame", "Stream'.«term_::_»"], - ["docBlame", "Stream'.«term_⊛_»"], - ["docBlame", "Stream'.«term_⋈_»"], ["docBlame", "Stream'.unfolds"], ["docBlame", "StrictWeakOrder.Equiv"], ["docBlame", "Submodule.«term_∙_»"], @@ -442,7 +417,6 @@ ["docBlame", "Filter.Germ.ofFun"], ["docBlame", "Filter.Realizer.F"], ["docBlame", "Filter.Realizer.σ"], - ["docBlame", "FirstOrder.Language.«term_≃ᴸ_»"], ["docBlame", "GenContFract.IntFractPair.b"], ["docBlame", "GenContFract.IntFractPair.fr"], ["docBlame", "GromovHausdorff.AuxGluingStruct.Space"], @@ -532,7 +506,6 @@ ["docBlame", "SimpleGraph.Subgraph.Adj"], ["docBlame", "SimpleGraph.Subgraph.verts"], ["docBlame", "SimpleGraph.Walk.notNilRec"], - ["docBlame", "Stream'.WSeq.«term_~ʷ_»"], ["docBlame", "Submodule.quotientPi_aux.invFun"], ["docBlame", "Submodule.quotientPi_aux.toFun"], ["docBlame", "Tactic.Elementwise.tacticElementwise!___"], From 464a3677e8b6545ffc1d10f0ed1c92e9c0a35f84 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Sat, 4 Jan 2025 14:04:14 +0000 Subject: [PATCH 137/189] chore: make uniqueness of conditionally complete linearly ordered archimedean fields independent of the construction of `Real` (#20242) Also avoid one import of `Analysis/` in `Algebra/`. Co-authored-by: Junyan Xu Co-authored-by: Junyan Xu --- Mathlib.lean | 1 + Mathlib/Algebra/Order/Archimedean/Basic.lean | 132 +++++++++++------- Mathlib/Algebra/Order/CompleteField.lean | 14 +- Mathlib/Algebra/Order/Field/Basic.lean | 18 +++ Mathlib/Algebra/Order/Ring/Abs.lean | 26 +++- .../Analysis/SpecialFunctions/Pow/Real.lean | 30 ---- Mathlib/Data/Real/CompleteField.lean | 22 +++ 7 files changed, 150 insertions(+), 93 deletions(-) create mode 100644 Mathlib/Data/Real/CompleteField.lean diff --git a/Mathlib.lean b/Mathlib.lean index 4ebb3fc8f8431..cf5693567c3ba 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2884,6 +2884,7 @@ import Mathlib.Data.Rat.Star import Mathlib.Data.Real.Archimedean import Mathlib.Data.Real.Basic import Mathlib.Data.Real.Cardinality +import Mathlib.Data.Real.CompleteField import Mathlib.Data.Real.ConjExponents import Mathlib.Data.Real.ENatENNReal import Mathlib.Data.Real.EReal diff --git a/Mathlib/Algebra/Order/Archimedean/Basic.lean b/Mathlib/Algebra/Order/Archimedean/Basic.lean index bf952ac20407b..d650574b73cde 100644 --- a/Mathlib/Algebra/Order/Archimedean/Basic.lean +++ b/Mathlib/Algebra/Order/Archimedean/Basic.lean @@ -303,11 +303,52 @@ theorem exists_nat_pow_near_of_lt_one (xpos : 0 < x) (hx : x ≤ 1) (ypos : 0 < end LinearOrderedSemifield section LinearOrderedField -variable [LinearOrderedField α] [Archimedean α] {x y ε : α} +variable [LinearOrderedField α] -theorem exists_rat_gt (x : α) : ∃ q : ℚ, x < q := - let ⟨n, h⟩ := exists_nat_gt x - ⟨n, by rwa [Rat.cast_natCast]⟩ +theorem archimedean_iff_nat_lt : Archimedean α ↔ ∀ x : α, ∃ n : ℕ, x < n := + ⟨@exists_nat_gt α _, fun H => + ⟨fun x y y0 => + (H (x / y)).imp fun n h => le_of_lt <| by rwa [div_lt_iff₀ y0, ← nsmul_eq_mul] at h⟩⟩ + +theorem archimedean_iff_nat_le : Archimedean α ↔ ∀ x : α, ∃ n : ℕ, x ≤ n := + archimedean_iff_nat_lt.trans + ⟨fun H x => (H x).imp fun _ => le_of_lt, fun H x => + let ⟨n, h⟩ := H x + ⟨n + 1, lt_of_le_of_lt h (Nat.cast_lt.2 (lt_add_one _))⟩⟩ + +theorem archimedean_iff_int_lt : Archimedean α ↔ ∀ x : α, ∃ n : ℤ, x < n := + ⟨@exists_int_gt α _, by + rw [archimedean_iff_nat_lt] + intro h x + obtain ⟨n, h⟩ := h x + refine ⟨n.toNat, h.trans_le ?_⟩ + exact mod_cast Int.self_le_toNat _⟩ + +theorem archimedean_iff_int_le : Archimedean α ↔ ∀ x : α, ∃ n : ℤ, x ≤ n := + archimedean_iff_int_lt.trans + ⟨fun H x => (H x).imp fun _ => le_of_lt, fun H x => + let ⟨n, h⟩ := H x + ⟨n + 1, lt_of_le_of_lt h (Int.cast_lt.2 (lt_add_one _))⟩⟩ + +theorem archimedean_iff_rat_lt : Archimedean α ↔ ∀ x : α, ∃ q : ℚ, x < q where + mp _ x := + let ⟨n, h⟩ := exists_nat_gt x + ⟨n, by rwa [Rat.cast_natCast]⟩ + mpr H := archimedean_iff_nat_lt.2 fun x ↦ + let ⟨q, h⟩ := H x; ⟨⌈q⌉₊, lt_of_lt_of_le h <| mod_cast Nat.le_ceil _⟩ + +theorem archimedean_iff_rat_le : Archimedean α ↔ ∀ x : α, ∃ q : ℚ, x ≤ q := + archimedean_iff_rat_lt.trans + ⟨fun H x => (H x).imp fun _ => le_of_lt, fun H x => + let ⟨n, h⟩ := H x + ⟨n + 1, lt_of_le_of_lt h (Rat.cast_lt.2 (lt_add_one _))⟩⟩ + +instance : Archimedean ℚ := + archimedean_iff_rat_le.2 fun q => ⟨q, by rw [Rat.cast_id]⟩ + +variable [Archimedean α] {x y ε : α} + +theorem exists_rat_gt (x : α) : ∃ q : ℚ, x < q := archimedean_iff_rat_lt.mp ‹_› _ theorem exists_rat_lt (x : α) : ∃ q : ℚ, (q : α) < x := let ⟨n, h⟩ := exists_int_lt x @@ -331,6 +372,44 @@ theorem exists_rat_btwn {x y : α} (h : x < y) : ∃ q : ℚ, x < q ∧ (q : α) subst H cases n0 +theorem exists_pow_btwn {n : ℕ} (hn : n ≠ 0) {x y : α} (h : x < y) (hy : 0 < y) : + ∃ q : α, 0 < q ∧ x < q ^ n ∧ q ^ n < y := by + have ⟨δ, δ_pos, cont⟩ := uniform_continuous_npow_on_bounded (max 1 y) + (sub_pos.mpr <| max_lt_iff.mpr ⟨h, hy⟩) n + have ex : ∃ m : ℕ, y ≤ (m * δ) ^ n := by + have ⟨m, hm⟩ := exists_nat_ge (y / δ + 1 / δ) + refine ⟨m, le_trans ?_ (le_self_pow₀ ?_ hn)⟩ <;> rw [← div_le_iff₀ δ_pos] + · exact (lt_add_of_pos_right _ <| by positivity).le.trans hm + · exact (le_add_of_nonneg_left <| by positivity).trans hm + let m := Nat.find ex + have m_pos : 0 < m := (Nat.find_pos _).mpr <| by simpa [zero_pow hn] using hy + let q := m.pred * δ + have qny : q ^ n < y := lt_of_not_le (Nat.find_min ex <| Nat.pred_lt m_pos.ne') + have q1y : |q| < max 1 y := (abs_eq_self.mpr <| by positivity).trans_lt <| lt_max_iff.mpr + (or_iff_not_imp_left.mpr fun q1 ↦ (le_self_pow₀ (le_of_not_lt q1) hn).trans_lt qny) + have xqn : max x 0 < q ^ n := + calc _ = y - (y - max x 0) := by rw [sub_sub_cancel] + _ ≤ (m * δ) ^ n - (y - max x 0) := sub_le_sub_right (Nat.find_spec ex) _ + _ < (m * δ) ^ n - ((m * δ) ^ n - q ^ n) := by + refine sub_lt_sub_left ((le_abs_self _).trans_lt <| cont _ _ q1y.le ?_) _ + rw [← Nat.succ_pred_eq_of_pos m_pos, Nat.cast_succ, ← sub_mul, + add_sub_cancel_left, one_mul, abs_eq_self.mpr (by positivity)] + _ = q ^ n := sub_sub_cancel .. + exact ⟨q, lt_of_le_of_ne (by positivity) fun q0 ↦ + (le_sup_right.trans_lt xqn).ne <| q0 ▸ (zero_pow hn).symm, le_sup_left.trans_lt xqn, qny⟩ + +@[deprecated (since := "2024-12-26")] alias exists_rat_pow_btwn_rat := exists_pow_btwn + +/-- There is a rational power between any two positive elements of an archimedean ordered field. -/ +theorem exists_rat_pow_btwn {n : ℕ} (hn : n ≠ 0) {x y : α} (h : x < y) (hy : 0 < y) : + ∃ q : ℚ, 0 < q ∧ x < (q : α) ^ n ∧ (q : α) ^ n < y := by + obtain ⟨q₂, hx₂, hy₂⟩ := exists_rat_btwn (max_lt h hy) + obtain ⟨q₁, hx₁, hq₁₂⟩ := exists_rat_btwn hx₂ + have : (0 : α) < q₂ := (le_max_right _ _).trans_lt hx₂ + norm_cast at hq₁₂ this + obtain ⟨q, hq, hq₁, hq₂⟩ := exists_pow_btwn hn hq₁₂ this + refine ⟨q, hq, (le_max_left _ _).trans_lt <| hx₁.trans ?_, hy₂.trans' ?_⟩ <;> assumption_mod_cast + theorem le_of_forall_rat_lt_imp_le (h : ∀ q : ℚ, (q : α) < x → (q : α) ≤ y) : x ≤ y := le_of_not_lt fun hyx => let ⟨_, hy, hx⟩ := exists_rat_btwn hyx @@ -365,48 +444,6 @@ theorem exists_rat_near (x : α) (ε0 : 0 < ε) : ∃ q : ℚ, |x - q| < ε := end LinearOrderedField -section LinearOrderedField - -variable [LinearOrderedField α] - -theorem archimedean_iff_nat_lt : Archimedean α ↔ ∀ x : α, ∃ n : ℕ, x < n := - ⟨@exists_nat_gt α _, fun H => - ⟨fun x y y0 => - (H (x / y)).imp fun n h => le_of_lt <| by rwa [div_lt_iff₀ y0, ← nsmul_eq_mul] at h⟩⟩ - -theorem archimedean_iff_nat_le : Archimedean α ↔ ∀ x : α, ∃ n : ℕ, x ≤ n := - archimedean_iff_nat_lt.trans - ⟨fun H x => (H x).imp fun _ => le_of_lt, fun H x => - let ⟨n, h⟩ := H x - ⟨n + 1, lt_of_le_of_lt h (Nat.cast_lt.2 (lt_add_one _))⟩⟩ - -theorem archimedean_iff_int_lt : Archimedean α ↔ ∀ x : α, ∃ n : ℤ, x < n := - ⟨@exists_int_gt α _, by - rw [archimedean_iff_nat_lt] - intro h x - obtain ⟨n, h⟩ := h x - refine ⟨n.toNat, h.trans_le ?_⟩ - exact mod_cast Int.self_le_toNat _⟩ - -theorem archimedean_iff_int_le : Archimedean α ↔ ∀ x : α, ∃ n : ℤ, x ≤ n := - archimedean_iff_int_lt.trans - ⟨fun H x => (H x).imp fun _ => le_of_lt, fun H x => - let ⟨n, h⟩ := H x - ⟨n + 1, lt_of_le_of_lt h (Int.cast_lt.2 (lt_add_one _))⟩⟩ - -theorem archimedean_iff_rat_lt : Archimedean α ↔ ∀ x : α, ∃ q : ℚ, x < q where - mp := @exists_rat_gt α _ - mpr H := archimedean_iff_nat_lt.2 fun x ↦ - let ⟨q, h⟩ := H x; ⟨⌈q⌉₊, lt_of_lt_of_le h <| mod_cast Nat.le_ceil _⟩ - -theorem archimedean_iff_rat_le : Archimedean α ↔ ∀ x : α, ∃ q : ℚ, x ≤ q := - archimedean_iff_rat_lt.trans - ⟨fun H x => (H x).imp fun _ => le_of_lt, fun H x => - let ⟨n, h⟩ := H x - ⟨n + 1, lt_of_le_of_lt h (Rat.cast_lt.2 (lt_add_one _))⟩⟩ - -end LinearOrderedField - instance : Archimedean ℕ := ⟨fun n m m0 => ⟨n, by rw [← mul_one n, smul_eq_mul, mul_assoc, one_mul m] @@ -419,9 +456,6 @@ instance : Archimedean ℤ := simpa only [nsmul_eq_mul, zero_add, mul_one] using mul_le_mul_of_nonneg_left (Int.add_one_le_iff.2 m0) (Int.ofNat_zero_le n.toNat)⟩⟩ -instance : Archimedean ℚ := - archimedean_iff_rat_le.2 fun q => ⟨q, by rw [Rat.cast_id]⟩ - instance Nonneg.instArchimedean [OrderedAddCommMonoid α] [Archimedean α] : Archimedean { x : α // 0 ≤ x } := ⟨fun x y hy => diff --git a/Mathlib/Algebra/Order/CompleteField.lean b/Mathlib/Algebra/Order/CompleteField.lean index 637abff4197d5..5ac0f4d4135a6 100644 --- a/Mathlib/Algebra/Order/CompleteField.lean +++ b/Mathlib/Algebra/Order/CompleteField.lean @@ -5,7 +5,6 @@ Authors: Alex J. Best, Yaël Dillies -/ import Mathlib.Algebra.Order.Archimedean.Hom import Mathlib.Algebra.Order.Group.Pointwise.CompleteLattice -import Mathlib.Analysis.SpecialFunctions.Pow.Real /-! # Conditionally complete linear ordered fields @@ -46,7 +45,7 @@ variable {F α β γ : Type*} noncomputable section -open Function Rat Real Set +open Function Rat Set open scoped Pointwise @@ -69,11 +68,6 @@ instance (priority := 100) ConditionallyCompleteLinearOrderedField.to_archimedea le_sub_iff_add_le.2 <| le_csSup _ _ ⟨x, forall_mem_range.2 h⟩ ⟨m+1, Nat.cast_succ m⟩) linarith) -/-- The reals are a conditionally complete linearly ordered field. -/ -instance : ConditionallyCompleteLinearOrderedField ℝ := - { (inferInstance : LinearOrderedField ℝ), - (inferInstance : ConditionallyCompleteLinearOrder ℝ) with } - namespace LinearOrderedField /-! @@ -329,10 +323,4 @@ theorem ringHom_monotone (hR : ∀ r : R, 0 ≤ r → ∃ s : R, s ^ 2 = r) (f : (monotone_iff_map_nonneg f).2 fun r h => by obtain ⟨s, rfl⟩ := hR r h; rw [map_pow]; apply sq_nonneg -/-- There exists no nontrivial ring homomorphism `ℝ →+* ℝ`. -/ -instance Real.RingHom.unique : Unique (ℝ →+* ℝ) where - default := RingHom.id ℝ - uniq f := congr_arg OrderRingHom.toRingHom (@Subsingleton.elim (ℝ →+*o ℝ) _ - ⟨f, ringHom_monotone (fun r hr => ⟨√r, sq_sqrt hr⟩) f⟩ default) - end Real diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index 029c8426a576b..e65d0da3e42fb 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -799,6 +799,24 @@ theorem abs_div (a b : α) : |a / b| = |a| / |b| := theorem abs_one_div (a : α) : |1 / a| = 1 / |a| := by rw [abs_div, abs_one] +theorem uniform_continuous_npow_on_bounded (B : α) {ε : α} (hε : 0 < ε) (n : ℕ) : + ∃ δ > 0, ∀ q r : α, |r| ≤ B → |q - r| ≤ δ → |q ^ n - r ^ n| < ε := by + wlog B_pos : 0 < B generalizing B + · have ⟨δ, δ_pos, cont⟩ := this 1 zero_lt_one + exact ⟨δ, δ_pos, fun q r hr ↦ cont q r (hr.trans ((le_of_not_lt B_pos).trans zero_le_one))⟩ + have pos : 0 < 1 + ↑n * (B + 1) ^ (n - 1) := zero_lt_one.trans_le <| le_add_of_nonneg_right <| + mul_nonneg n.cast_nonneg <| (pow_pos (B_pos.trans <| lt_add_of_pos_right _ zero_lt_one) _).le + refine ⟨min 1 (ε / (1 + n * (B + 1) ^ (n - 1))), lt_min zero_lt_one (div_pos hε pos), + fun q r hr hqr ↦ (abs_pow_sub_pow_le ..).trans_lt ?_⟩ + rw [le_inf_iff, le_div_iff₀ pos, mul_one_add, ← mul_assoc] at hqr + obtain h | h := (abs_nonneg (q - r)).eq_or_lt + · simpa only [← h, zero_mul] using hε + refine (lt_of_le_of_lt ?_ <| lt_add_of_pos_left _ h).trans_le hqr.2 + refine mul_le_mul_of_nonneg_left (pow_le_pow_left₀ ((abs_nonneg _).trans le_sup_left) ?_ _) + (mul_nonneg (abs_nonneg _) n.cast_nonneg) + refine max_le ?_ (hr.trans <| le_add_of_nonneg_right zero_le_one) + exact add_sub_cancel r q ▸ (abs_add_le ..).trans (add_le_add hr hqr.1) + end namespace Mathlib.Meta.Positivity diff --git a/Mathlib/Algebra/Order/Ring/Abs.lean b/Mathlib/Algebra/Order/Ring/Abs.lean index 68c3cfc6d55b0..fa980a32dacce 100644 --- a/Mathlib/Algebra/Order/Ring/Abs.lean +++ b/Mathlib/Algebra/Order/Ring/Abs.lean @@ -143,7 +143,7 @@ end LinearOrderedRing section LinearOrderedCommRing -variable [LinearOrderedCommRing α] +variable [LinearOrderedCommRing α] (a b : α) (n : ℕ) theorem abs_sub_sq (a b : α) : |a - b| * |a - b| = a * a + b * b - (1 + 1) * a * b := by rw [abs_mul_abs_self] @@ -153,6 +153,30 @@ theorem abs_sub_sq (a b : α) : |a - b| * |a - b| = a * a + b * b - (1 + 1) * a lemma abs_unit_intCast (a : ℤˣ) : |((a : ℤ) : α)| = 1 := by cases Int.units_eq_one_or a <;> simp_all +private def geomSum : ℕ → α + | 0 => 1 + | n + 1 => a * geomSum n + b ^ (n + 1) + +private theorem abs_geomSum_le : |geomSum a b n| ≤ (n + 1) * max |a| |b| ^ n := by + induction' n with n ih; · simp [geomSum] + refine (abs_add_le ..).trans ?_ + rw [abs_mul, abs_pow, Nat.cast_succ, add_one_mul] + refine add_le_add ?_ (pow_le_pow_left₀ (abs_nonneg _) le_sup_right _) + rw [pow_succ, ← mul_assoc, mul_comm |a|] + exact mul_le_mul ih le_sup_left (abs_nonneg _) (mul_nonneg + (@Nat.cast_succ α .. ▸ Nat.cast_nonneg _) <| pow_nonneg ((abs_nonneg _).trans le_sup_left) _) + +private theorem pow_sub_pow_eq_sub_mul_geomSum : + a ^ (n + 1) - b ^ (n + 1) = (a - b) * geomSum a b n := by + induction' n with n ih; · simp [geomSum] + rw [geomSum, mul_add, mul_comm a, ← mul_assoc, ← ih, + sub_mul, sub_mul, ← pow_succ, ← pow_succ', mul_comm, sub_add_sub_cancel] + +theorem abs_pow_sub_pow_le : |a ^ n - b ^ n| ≤ |a - b| * n * max |a| |b| ^ (n - 1) := by + obtain _ | n := n; · simp + rw [Nat.add_sub_cancel, pow_sub_pow_eq_sub_mul_geomSum, abs_mul, mul_assoc, Nat.cast_succ] + exact mul_le_mul_of_nonneg_left (abs_geomSum_le ..) (abs_nonneg _) + end LinearOrderedCommRing section diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index 809b94d5aa7d0..7fa60d874f8d7 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -970,36 +970,6 @@ theorem rpow_div_two_eq_sqrt {x : ℝ} (r : ℝ) (hx : 0 ≤ x) : x ^ (r / 2) = end Sqrt -variable {n : ℕ} - -theorem exists_rat_pow_btwn_rat_aux (hn : n ≠ 0) (x y : ℝ) (h : x < y) (hy : 0 < y) : - ∃ q : ℚ, 0 < q ∧ x < (q : ℝ) ^ n ∧ (q : ℝ) ^ n < y := by - have hn' : 0 < (n : ℝ) := mod_cast hn.bot_lt - obtain ⟨q, hxq, hqy⟩ := - exists_rat_btwn (rpow_lt_rpow (le_max_left 0 x) (max_lt hy h) <| inv_pos.mpr hn') - have := rpow_nonneg (le_max_left 0 x) n⁻¹ - have hq := this.trans_lt hxq - replace hxq := rpow_lt_rpow this hxq hn' - replace hqy := rpow_lt_rpow hq.le hqy hn' - rw [rpow_natCast, rpow_natCast, rpow_inv_natCast_pow _ hn] at hxq hqy - · exact ⟨q, mod_cast hq, (le_max_right _ _).trans_lt hxq, hqy⟩ - · exact hy.le - · exact le_max_left _ _ - -theorem exists_rat_pow_btwn_rat (hn : n ≠ 0) {x y : ℚ} (h : x < y) (hy : 0 < y) : - ∃ q : ℚ, 0 < q ∧ x < q ^ n ∧ q ^ n < y := by - apply_mod_cast exists_rat_pow_btwn_rat_aux hn x y <;> assumption - -/-- There is a rational power between any two positive elements of an archimedean ordered field. -/ -theorem exists_rat_pow_btwn {α : Type*} [LinearOrderedField α] [Archimedean α] (hn : n ≠ 0) - {x y : α} (h : x < y) (hy : 0 < y) : ∃ q : ℚ, 0 < q ∧ x < (q : α) ^ n ∧ (q : α) ^ n < y := by - obtain ⟨q₂, hx₂, hy₂⟩ := exists_rat_btwn (max_lt h hy) - obtain ⟨q₁, hx₁, hq₁₂⟩ := exists_rat_btwn hx₂ - have : (0 : α) < q₂ := (le_max_right _ _).trans_lt hx₂ - norm_cast at hq₁₂ this - obtain ⟨q, hq, hq₁, hq₂⟩ := exists_rat_pow_btwn_rat hn hq₁₂ this - refine ⟨q, hq, (le_max_left _ _).trans_lt <| hx₁.trans ?_, hy₂.trans' ?_⟩ <;> assumption_mod_cast - end Real namespace Complex diff --git a/Mathlib/Data/Real/CompleteField.lean b/Mathlib/Data/Real/CompleteField.lean new file mode 100644 index 0000000000000..91bedb3dfe86e --- /dev/null +++ b/Mathlib/Data/Real/CompleteField.lean @@ -0,0 +1,22 @@ +/- +Copyright (c) 2024 Alex J. Best. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Alex J. Best +-/ +import Mathlib.Algebra.Order.CompleteField +import Mathlib.Data.Real.Sqrt + +/-! +# The reals are a conditionally complete linearly ordered field +-/ + +/-- The reals are a conditionally complete linearly ordered field. -/ +noncomputable instance : ConditionallyCompleteLinearOrderedField ℝ := + { (inferInstance : LinearOrderedField ℝ), + (inferInstance : ConditionallyCompleteLinearOrder ℝ) with } + +/-- There exists no nontrivial ring homomorphism `ℝ →+* ℝ`. -/ +instance Real.RingHom.unique : Unique (ℝ →+* ℝ) where + default := RingHom.id ℝ + uniq f := congr_arg OrderRingHom.toRingHom (@Subsingleton.elim (ℝ →+*o ℝ) _ + ⟨f, ringHom_monotone (fun r hr => ⟨√r, sq_sqrt hr⟩) f⟩ default) From 6d2cbb58e68c7424d746005d5bbf75a31cb14cb6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 4 Jan 2025 15:06:22 +0000 Subject: [PATCH 138/189] feat: the empty set is a topological basis iff the space is empty (#20441) --- Mathlib/Topology/Bases.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/Topology/Bases.lean b/Mathlib/Topology/Bases.lean index 969cd3b6dd3fa..691e389d9e045 100644 --- a/Mathlib/Topology/Bases.lean +++ b/Mathlib/Topology/Bases.lean @@ -290,6 +290,10 @@ protected theorem IsTopologicalBasis.continuous_iff {β : Type*} [TopologicalSpa Continuous f ↔ ∀ s ∈ B, IsOpen (f ⁻¹' s) := by rw [hB.eq_generateFrom, continuous_generateFrom_iff] +@[simp] lemma isTopologicalBasis_empty : IsTopologicalBasis (∅ : Set (Set α)) ↔ IsEmpty α where + mp h := by simpa using h.sUnion_eq.symm + mpr h := ⟨by simp, by simp [Set.univ_eq_empty_iff.2], Subsingleton.elim ..⟩ + variable (α) /-- A separable space is one with a countable dense subset, available through From 0061a577b405448e5bd709c059c2ac246fc0b1d3 Mon Sep 17 00:00:00 2001 From: Nailin Guan <150537269+Thmoas-Guan@users.noreply.github.com> Date: Sat, 4 Jan 2025 15:57:11 +0000 Subject: [PATCH 139/189] feat(Topology/Group): Lemmas about profinite group (#20282) A lemma stating a topological group that has a `ContinuousMulEquiv` to a profinite group is profinite. With addition of a lemma `Homeomorph.totallyDisconnectedSpace` --- Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean | 7 +++++++ Mathlib/Topology/Homeomorph.lean | 5 +++++ 2 files changed, 12 insertions(+) diff --git a/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean b/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean index 2e6184e78b0ae..c61566db4fee5 100644 --- a/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean +++ b/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean @@ -167,6 +167,13 @@ def ofClosedSubgroup {G : ProfiniteGrp} (H : ClosedSubgroup G) : ProfiniteGrp : letI : CompactSpace H := inferInstance of H.1 +/-- A topological group that has a `ContinuousMulEquiv` to a profinite group is profinite. -/ +def ofContinuousMulEquiv {G : ProfiniteGrp.{u}} {H : Type v} [TopologicalSpace H] + [Group H] [TopologicalGroup H] (e : G ≃ₜ* H) : ProfiniteGrp.{v} := + let _ : CompactSpace H := Homeomorph.compactSpace e.toHomeomorph + let _ : TotallyDisconnectedSpace H := Homeomorph.totallyDisconnectedSpace e.toHomeomorph + .of H + /-- The functor mapping a profinite group to its underlying profinite space. -/ def profiniteGrpToProfinite : ProfiniteGrp ⥤ Profinite where obj G := G.toProfinite diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index 224fd0408685c..52f9fb7edaad3 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -313,6 +313,11 @@ protected theorem t3Space [T3Space X] (h : X ≃ₜ Y) : T3Space Y := h.symm.isE theorem isDenseEmbedding (h : X ≃ₜ Y) : IsDenseEmbedding h := { h.isEmbedding with dense := h.surjective.denseRange } +protected lemma totallyDisconnectedSpace (h : X ≃ₜ Y) [tdc : TotallyDisconnectedSpace X] : + TotallyDisconnectedSpace Y := + (totallyDisconnectedSpace_iff Y).mpr + (h.range_coe ▸ ((IsEmbedding.isTotallyDisconnected_range h.isEmbedding).mpr tdc)) + @[deprecated (since := "2024-09-30")] alias denseEmbedding := isDenseEmbedding From a36b4f2f1dc49c3c40c95b80d96fbba3f1ec98f1 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Sat, 4 Jan 2025 18:31:41 +0000 Subject: [PATCH 140/189] chore(RingTheory/Generators): make algebra instance a def (#14265) Makes the algebra instance `Algebra P.Ring S` for `P : Generators R S` a def and adds the instance attribute locally where needed. This prevents an instance diamond in the case where `S` is a quotient of some polynomial algebra. See https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Diamond.20in.20.60Algebra.2EGenerators.60. Co-authored-by: Andrew Yang --- Mathlib/RingTheory/Generators.lean | 31 ++++++++++++------- Mathlib/RingTheory/Kaehler/JacobiZariski.lean | 6 ++-- Mathlib/RingTheory/Presentation.lean | 9 +++--- Mathlib/RingTheory/Smooth/StandardSmooth.lean | 4 ++- 4 files changed, 31 insertions(+), 19 deletions(-) diff --git a/Mathlib/RingTheory/Generators.lean b/Mathlib/RingTheory/Generators.lean index d63cf6d23df0c..417614c216301 100644 --- a/Mathlib/RingTheory/Generators.lean +++ b/Mathlib/RingTheory/Generators.lean @@ -53,9 +53,16 @@ structure Algebra.Generators where /-- A section of `R[X] → S`. -/ σ' : S → MvPolynomial vars R aeval_val_σ' : ∀ s, aeval val (σ' s) = s + /-- An `R[X]`-algebra instance on `S`. The default is the one induced by the map `R[X] → S`, + but this causes a diamond if there is an existing instance. -/ + algebra : Algebra (MvPolynomial vars R) S := (aeval val).toAlgebra + algebraMap_eq : + algebraMap (MvPolynomial vars R) S = aeval (R := R) val := by rfl namespace Algebra.Generators +attribute [instance] algebra + variable {R S} variable (P : Generators.{w} R S) @@ -74,16 +81,13 @@ initialize_simps_projections Algebra.Generators (σ' → σ) @[simp] lemma aeval_val_σ (s) : aeval P.val (P.σ s) = s := P.aeval_val_σ' s -instance : Algebra P.Ring S := (aeval P.val).toAlgebra - noncomputable instance {R₀} [CommRing R₀] [Algebra R₀ R] [Algebra R₀ S] [IsScalarTower R₀ R S] : - IsScalarTower R₀ P.Ring S := IsScalarTower.of_algebraMap_eq' - ((aeval (R := R) P.val).comp_algebraMap_of_tower R₀).symm - -lemma algebraMap_eq : algebraMap P.Ring S = ↑(aeval (R := R) P.val) := rfl + IsScalarTower R₀ P.Ring S := IsScalarTower.of_algebraMap_eq' <| + P.algebraMap_eq ▸ ((aeval (R := R) P.val).comp_algebraMap_of_tower R₀).symm @[simp] -lemma algebraMap_apply (x) : algebraMap P.Ring S x = aeval (R := R) P.val x := rfl +lemma algebraMap_apply (x) : algebraMap P.Ring S x = aeval (R := R) P.val x := by + simp [algebraMap_eq] @[simp] lemma σ_smul (x y) : P.σ x • y = x * y := by @@ -93,7 +97,8 @@ lemma σ_injective : P.σ.Injective := by intro x y e rw [← P.aeval_val_σ x, ← P.aeval_val_σ y, e] -lemma algebraMap_surjective : Function.Surjective (algebraMap P.Ring S) := (⟨_, P.aeval_val_σ ·⟩) +lemma algebraMap_surjective : Function.Surjective (algebraMap P.Ring S) := + (⟨_, P.algebraMap_apply _ ▸ P.aeval_val_σ ·⟩) section Construction @@ -148,7 +153,7 @@ noncomputable def toExtension : Extension R S where Ring := P.Ring σ := P.σ - algebraMap_σ := P.aeval_val_σ + algebraMap_σ := by simp section Localization @@ -386,7 +391,7 @@ lemma ofComp_toAlgHom_monomial_sumElim (Q : Generators S T) (P : Generators R S) simp only [MvPolynomial.algebraMap_apply, ofComp_val, aeval_monomial] rw [Finsupp.prod_sumElim] simp only [Function.comp_def, Sum.elim_inl, Sum.elim_inr, ← map_pow, ← map_finsupp_prod, - C_mul, algebra.smul_def, MvPolynomial.algebraMap_apply, mul_assoc] + C_mul, Algebra.smul_def, MvPolynomial.algebraMap_apply, mul_assoc] nth_rw 2 [mul_comm] lemma toComp_toAlgHom_monomial (Q : Generators S T) (P : Generators R S) (j a) : @@ -427,9 +432,13 @@ lemma Hom.toExtensionHom_comp [Algebra R S'] [IsScalarTower R S S'] /-- The kernel of a presentation. -/ noncomputable abbrev ker : Ideal P.Ring := P.toExtension.ker -lemma ker_eq_ker_aeval_val : P.ker = RingHom.ker (aeval P.val) := +lemma ker_eq_ker_aeval_val : P.ker = RingHom.ker (aeval P.val) := by + simp only [ker, Extension.ker, toExtension_Ring, algebraMap_eq] rfl +variable {P} in +lemma aeval_val_eq_zero {x} (hx : x ∈ P.ker) : aeval P.val x = 0 := by rwa [← algebraMap_apply] + lemma map_toComp_ker (Q : Generators S T) (P : Generators R S) : P.ker.map (Q.toComp P).toAlgHom.toRingHom = RingHom.ker (Q.ofComp P).toAlgHom := by letI : DecidableEq (Q.vars →₀ ℕ) := Classical.decEq _ diff --git a/Mathlib/RingTheory/Kaehler/JacobiZariski.lean b/Mathlib/RingTheory/Kaehler/JacobiZariski.lean index c789134523f00..dba3d8955d251 100644 --- a/Mathlib/RingTheory/Kaehler/JacobiZariski.lean +++ b/Mathlib/RingTheory/Kaehler/JacobiZariski.lean @@ -54,7 +54,7 @@ def kerCompPreimage (x : Q.ker) : refine rename ?_ (P.σ r) * monomial ?_ 1 exacts [Sum.inr, n.mapDomain Sum.inl] · simp only [ker_eq_ker_aeval_val, RingHom.mem_ker] - conv_rhs => rw [← show aeval Q.val x.1 = 0 from x.2, ← x.1.support_sum_monomial_coeff] + conv_rhs => rw [← aeval_val_eq_zero x.2, ← x.1.support_sum_monomial_coeff] simp only [Finsupp.sum, map_sum, map_mul, aeval_rename, Function.comp_def, comp_val, Sum.elim_inr, aeval_monomial, map_one, Finsupp.prod_mapDomain_index_inj Sum.inl_injective, Sum.elim_inl, one_mul] @@ -120,7 +120,7 @@ lemma Cotangent.exact : refine MvPolynomial.algHom_ext fun i ↦ ?_ show (Q.ofComp P).toAlgHom ((Q.toComp P).toAlgHom (X i)) = _ simp - · simp [-self_vars, show aeval P.val x = 0 from hx] + · simp [-self_vars, aeval_val_eq_zero hx] · intro x hx obtain ⟨⟨x : (Q.comp P).Ring, hx'⟩, rfl⟩ := Extension.Cotangent.mk_surjective x replace hx : (Q.ofComp P).toAlgHom x ∈ Q.ker ^ 2 := by @@ -149,7 +149,7 @@ lemma Cotangent.exact : toExtension_Ring, toExtension_commRing, toExtension_algebra₂] refine ⟨?_, Submodule.subset_span ⟨Extension.Cotangent.mk ⟨w, hw⟩, ?_⟩⟩ · simp only [ker_eq_ker_aeval_val, RingHom.mem_ker, Hom.algebraMap_toAlgHom] - rw [show aeval P.val w = 0 from hw, map_zero] + rw [aeval_val_eq_zero hw, map_zero] · rw [map_mk] rfl diff --git a/Mathlib/RingTheory/Presentation.lean b/Mathlib/RingTheory/Presentation.lean index 68fbb2cf28452..1cca0b2f225c8 100644 --- a/Mathlib/RingTheory/Presentation.lean +++ b/Mathlib/RingTheory/Presentation.lean @@ -78,7 +78,8 @@ protected abbrev Quotient : Type (max w u) := P.Ring ⧸ P.ker /-- `P.Quotient` is `P.Ring`-isomorphic to `S` and in particular `R`-isomorphic to `S`. -/ def quotientEquiv : P.Quotient ≃ₐ[P.Ring] S := - Ideal.quotientKerAlgEquivOfRightInverse (f := Algebra.ofId P.Ring S) P.aeval_val_σ + Ideal.quotientKerAlgEquivOfRightInverse (f := Algebra.ofId P.Ring S) (g := P.σ) <| fun x ↦ by + rw [Algebra.ofId_apply, P.algebraMap_apply, P.aeval_val_σ] @[simp] lemma quotientEquiv_mk (p : P.Ring) : P.quotientEquiv p = algebraMap P.Ring S p := @@ -393,9 +394,9 @@ private lemma span_range_relation_eq_ker_comp : Ideal.span (Set.range (Sum.elim (Algebra.Presentation.comp_relation_aux Q P) fun rp ↦ (rename Sum.inr) (P.relation rp))) = (Q.comp P.toGenerators).ker := by rw [Generators.ker_eq_ker_aeval_val, Q.aeval_comp_val_eq, ← AlgHom.comap_ker] - show _ = Ideal.comap _ (Q.ker) - rw [← Q.span_range_relation_eq_ker, ← Q.aux_image_relation P, ← Ideal.map_span, - Ideal.comap_map_of_surjective' _ (Q.aux_surjective P)] + show _ = Ideal.comap _ (RingHom.ker (aeval Q.val)) + rw [← Q.ker_eq_ker_aeval_val, ← Q.span_range_relation_eq_ker, ← Q.aux_image_relation P, + ← Ideal.map_span, Ideal.comap_map_of_surjective' _ (Q.aux_surjective P)] rw [Set.Sum.elim_range, Ideal.span_union, Q.aux_ker, ← P.ker_eq_ker_aeval_val, ← P.span_range_relation_eq_ker, Ideal.map_span] congr diff --git a/Mathlib/RingTheory/Smooth/StandardSmooth.lean b/Mathlib/RingTheory/Smooth/StandardSmooth.lean index 9aef55d125b25..5efb276bcdc27 100644 --- a/Mathlib/RingTheory/Smooth/StandardSmooth.lean +++ b/Mathlib/RingTheory/Smooth/StandardSmooth.lean @@ -295,7 +295,8 @@ private lemma jacobiMatrix_comp_₁₁_det : rw [jacobian_eq_jacobiMatrix_det, AlgHom.map_det (aeval (Q.comp P).val), RingHom.map_det] congr ext i j : 1 - simp only [Matrix.map_apply, RingHom.mapMatrix_apply, ← Q.jacobiMatrix_comp_inl_inl P] + simp only [Matrix.map_apply, RingHom.mapMatrix_apply, ← Q.jacobiMatrix_comp_inl_inl P, + Q.algebraMap_apply] apply aeval_sum_elim end Q @@ -383,6 +384,7 @@ lemma baseChange_jacobian : (P.baseChange T).jacobian = 1 ⊗ₜ P.jacobian := b rfl rw [h] erw [← RingHom.map_det, aeval_map_algebraMap] + rw [P.algebraMap_apply] apply aeval_one_tmul end BaseChange From 83411aa397325c50ebaa2c0344f02140e58496f7 Mon Sep 17 00:00:00 2001 From: Lucas Whitfield Date: Sat, 4 Jan 2025 18:40:12 +0000 Subject: [PATCH 141/189] feat(Algebra/Lie): add Lie's theorem (#13480) Prove Lie's theorem, that Lie modules of solvable Lie algebras over algebraically closed fields of characteristic 0 have a common eigenvector for the action of all elements of the Lie algebra. --- Mathlib.lean | 1 + Mathlib/Algebra/Lie/LieTheorem.lean | 250 ++++++++++++++++++++++++ Mathlib/Algebra/Lie/Submodule.lean | 4 + Mathlib/Algebra/Lie/TraceForm.lean | 3 +- Mathlib/Algebra/Lie/Weights/Linear.lean | 17 +- Mathlib/LinearAlgebra/Trace.lean | 9 +- Mathlib/RingTheory/Trace/Basic.lean | 10 +- 7 files changed, 280 insertions(+), 14 deletions(-) create mode 100644 Mathlib/Algebra/Lie/LieTheorem.lean diff --git a/Mathlib.lean b/Mathlib.lean index cf5693567c3ba..9ccfe90d0fd2c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -503,6 +503,7 @@ import Mathlib.Algebra.Lie.Free import Mathlib.Algebra.Lie.IdealOperations import Mathlib.Algebra.Lie.InvariantForm import Mathlib.Algebra.Lie.Killing +import Mathlib.Algebra.Lie.LieTheorem import Mathlib.Algebra.Lie.Matrix import Mathlib.Algebra.Lie.Nilpotent import Mathlib.Algebra.Lie.NonUnitalNonAssocAlgebra diff --git a/Mathlib/Algebra/Lie/LieTheorem.lean b/Mathlib/Algebra/Lie/LieTheorem.lean new file mode 100644 index 0000000000000..a9e21ea679bd7 --- /dev/null +++ b/Mathlib/Algebra/Lie/LieTheorem.lean @@ -0,0 +1,250 @@ +/- +Copyright (c) 2024 Lucas Whitfield. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Lucas Whitfield, Johan Commelin +-/ +import Mathlib.Algebra.Lie.Weights.Basic +import Mathlib.RingTheory.Finiteness.Nilpotent + +/-! +# Lie's theorem for Solvable Lie algebras. + +Lie's theorem asserts that Lie modules of solvable Lie algebras over fields of characteristic 0 +have a common eigenvector for the action of all elements of the Lie algebra. +This result is named `LieModule.exists_forall_lie_eq_smul_of_isSolvable`. +-/ + +namespace LieModule + +section + +/- +The following variables generalize the setting where: +- `R` is a principal ideal domain of characteristic zero, +- `L` is a Lie algebra over `R`, +- `V` is a Lie algebra module over `L` +- `A` is a Lie ideal of `L`. +Besides generalizing, it also make the proof of `lie_stable` syntactically smoother. +-/ +variable {R L A V : Type*} [CommRing R] +variable [IsPrincipalIdealRing R] [IsDomain R] [CharZero R] +variable [LieRing L] [LieAlgebra R L] +variable [LieRing A] [LieAlgebra R A] +variable [Bracket L A] [Bracket A L] +variable [AddCommGroup V] [Module R V] [Module.Free R V] [Module.Finite R V] +variable [LieRingModule L V] [LieModule R L V] +variable [LieRingModule A V] [LieModule R A V] +variable [IsLieTower L A V] [IsLieTower A L V] + +variable (χ : A → R) + +open Module (finrank) +open LieModule + +local notation "π" => LieModule.toEnd R _ V + +private abbrev T (w : A) : Module.End R V := (π w) - χ w • 1 + +/-- An auxiliary lemma used only in the definition `LieModule.weightSpaceOfIsLieTower` below. -/ +private lemma weightSpaceOfIsLieTower_aux (z : L) (v : V) (hv : v ∈ weightSpace V χ) : + ⁅z, v⁆ ∈ weightSpace V χ := by + rw [mem_weightSpace] at hv ⊢ + intro a + rcases eq_or_ne v 0 with (rfl | hv') + · simp only [lie_zero, smul_zero] + suffices χ ⁅z, a⁆ = 0 by + rw [leibniz_lie, hv a, lie_smul, lie_swap_lie, hv, this, zero_smul, neg_zero, zero_add] + let U' : ℕ →o Submodule R V := + { toFun n := Submodule.span R {((π z)^i) v | i < n}, + monotone' i j h := Submodule.span_mono (fun _ ⟨c, hc, hw⟩ ↦ ⟨c, lt_of_lt_of_le hc h, hw⟩) } + have map_U'_le (n : ℕ) : Submodule.map (π z) (U' n) ≤ U' (n + 1) := by + simp only [OrderHom.coe_mk, Submodule.map_span, toEnd_apply_apply, U'] + apply Submodule.span_mono + suffices ∀ a < n, ∃ b < n + 1, ((π z) ^ b) v = ((π z) ^ (a + 1)) v by simpa [pow_succ'] + aesop + have T_apply_succ (w : A) (n : ℕ) : + Submodule.map (T χ w) (U' (n + 1)) ≤ U' n := by + simp only [OrderHom.coe_mk, U', Submodule.map_span, Submodule.span_le, Set.image_subset_iff] + simp only [Set.subset_def, Set.mem_setOf_eq, Set.mem_preimage, SetLike.mem_coe, + forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] + induction n generalizing w + · simp only [zero_add, Nat.lt_one_iff, LinearMap.sub_apply, LieModule.toEnd_apply_apply, + LinearMap.smul_apply, LinearMap.one_apply, forall_eq, pow_zero, hv w, sub_self, zero_mem] + · next n hn => + intro m hm + obtain (hm | rfl) : m < n + 1 ∨ m = n + 1 := by omega + · exact U'.mono (Nat.le_succ n) (hn w m hm) + have H : ∀ w, ⁅w, (π z ^ n) v⁆ = (T χ w) ((π z ^ n) v) + χ w • ((π z ^ n) v) := by simp + rw [T, LinearMap.sub_apply, pow_succ', LinearMap.mul_apply, LieModule.toEnd_apply_apply, + LieModule.toEnd_apply_apply, LinearMap.smul_apply, LinearMap.one_apply, leibniz_lie, + lie_swap_lie w z, H, H, lie_add, lie_smul, add_sub_assoc, add_sub_assoc, sub_self, add_zero] + refine add_mem (neg_mem <| add_mem ?_ ?_) ?_ + · exact U'.mono n.le_succ (hn _ n n.lt_succ_self) + · exact Submodule.smul_mem _ _ (Submodule.subset_span ⟨n, n.lt_succ_self, rfl⟩) + · exact map_U'_le _ <| Submodule.mem_map_of_mem <| hn w n n.lt_succ_self + set U : LieSubmodule R A V := + { toSubmodule := ⨆ k : ℕ, U' k + lie_mem {w} x hx := by + rw [show ⁅w, x⁆ = (T χ w) x + χ w • x by simp] + apply add_mem _ (Submodule.smul_mem _ _ hx) + set U := ⨆ k : ℕ, U' k + suffices Submodule.map (T χ w) U ≤ U from this <| Submodule.mem_map_of_mem hx + rw [Submodule.map_iSup, iSup_le_iff] + rintro (_|i) + · simp [U', Submodule.map_span] + · exact (T_apply_succ w i).trans (le_iSup _ _) } + have hzU (x : V) (hx : x ∈ U) : (π z) x ∈ U := by + suffices Submodule.map (π z) U ≤ U from this <| Submodule.mem_map_of_mem hx + simp only [U, Submodule.map_iSup, iSup_le_iff] + exact fun i ↦ (map_U'_le i).trans (le_iSup _ _) + have trace_za_zero : (LieModule.toEnd R A _ ⁅z, a⁆).trace R U = 0 := by + have hres : LieModule.toEnd R A U ⁅z, a⁆ = ⁅(π z).restrict hzU, LieModule.toEnd R A U a⁆ := by + ext ⟨x, hx⟩ + show ⁅⁅z, a⁆, x⁆ = ⁅z, ⁅a, x⁆⁆ - ⁅a, ⁅z, x⁆⁆ + simp only [leibniz_lie z a, add_sub_cancel_right] + rw [hres, LinearMap.trace_lie] + have trace_T_U_zero (w : A) : (T χ w).trace R U = 0 := by + have key (i : ℕ) (hi : i ≠ 0) : ∃ j < i, Submodule.map (T χ w) (U' i) ≤ U' j := by + obtain ⟨j, rfl⟩ := Nat.exists_eq_succ_of_ne_zero hi + exact ⟨j, j.lt_succ_self, T_apply_succ w j⟩ + apply IsNilpotent.eq_zero + apply LinearMap.isNilpotent_trace_of_isNilpotent + rw [Module.Finite.Module.End.isNilpotent_iff_of_finite] + suffices ⨆ i, U' i ≤ Module.End.maxGenEigenspace (T χ w) 0 by + intro x + specialize this x.2 + simp only [Module.End.mem_maxGenEigenspace, zero_smul, sub_zero] at this + peel this with n hn + ext + simp only [ZeroMemClass.coe_zero, ← hn]; clear hn + induction n <;> simp_all [pow_succ'] + apply iSup_le + intro i x hx + simp only [Module.End.mem_maxGenEigenspace, zero_smul, sub_zero] + induction i using Nat.strong_induction_on generalizing x + next i ih => + obtain rfl | hi := eq_or_ne i 0 + · simp_all [U'] + obtain ⟨j, hj, hj'⟩ := key i hi + obtain ⟨k, hk⟩ := ih j hj (hj' <| Submodule.mem_map_of_mem hx) + use k+1 + rw [pow_succ, LinearMap.mul_apply, hk] + have trace_za : (toEnd R A _ ⁅z, a⁆).trace R U = χ ⁅z, a⁆ • (finrank R U) := by + simpa [T, sub_eq_zero] using trace_T_U_zero ⁅z, a⁆ + suffices finrank R U ≠ 0 by simp_all + suffices Nontrivial U from Module.finrank_pos.ne' + have hvU : v ∈ U := by + apply Submodule.mem_iSup_of_mem 1 + apply Submodule.subset_span + use 0, zero_lt_one + rw [pow_zero, LinearMap.one_apply] + exact nontrivial_of_ne ⟨v, hvU⟩ 0 <| by simp [hv'] + +variable (R V) in +/-- The weight space of `V` with respect to `χ : A → R`, a priori a Lie submodule for `A`, is also a +Lie submodule for `L`. -/ +def weightSpaceOfIsLieTower (χ : A → R) : LieSubmodule R L V := + { toSubmodule := weightSpace V χ + lie_mem {z v} hv := weightSpaceOfIsLieTower_aux χ z v hv } + +end + +section + +variable {k : Type*} [Field k] +variable {L : Type*} [LieRing L] [LieAlgebra k L] +variable {V : Type*} [AddCommGroup V] [Module k V] [LieRingModule L V] [LieModule k L V] + +variable [CharZero k] [Module.Finite k V] + +open Submodule in +theorem exists_nontrivial_weightSpace_of_lieIdeal [LieModule.IsTriangularizable k L V] + (A : LieIdeal k L) (hA : IsCoatom A.toSubmodule) + (χ₀ : Module.Dual k A) [Nontrivial (weightSpace V χ₀)] : + ∃ (χ : Module.Dual k L), Nontrivial (weightSpace V χ) := by + obtain ⟨z, -, hz⟩ := SetLike.exists_of_lt (hA.lt_top) + let e : (k ∙ z) ≃ₗ[k] k := (LinearEquiv.toSpanNonzeroSingleton k L z <| by aesop).symm + have he : ∀ x, e x • z = x := by simp [e] + have hA : IsCompl A.toSubmodule (k ∙ z) := isCompl_span_singleton_of_isCoatom_of_not_mem hA hz + let π₁ : L →ₗ[k] A := A.toSubmodule.linearProjOfIsCompl (k ∙ z) hA + let π₂ : L →ₗ[k] (k ∙ z) := (k ∙ z).linearProjOfIsCompl ↑A hA.symm + + set W : LieSubmodule k L V := weightSpaceOfIsLieTower k V χ₀ + obtain ⟨c, hc⟩ : ∃ c, (toEnd k _ W z).HasEigenvalue c := by + have : Nontrivial W := inferInstanceAs (Nontrivial (weightSpace V χ₀)) + apply Module.End.exists_hasEigenvalue_of_genEigenspace_eq_top + exact LieModule.IsTriangularizable.maxGenEigenspace_eq_top z + + obtain ⟨⟨v, hv⟩, hvc⟩ := hc.exists_hasEigenvector + have hv' : ∀ (x : ↥A), ⁅x, v⁆ = χ₀ x • v := by + simpa [W, weightSpaceOfIsLieTower, mem_weightSpace] using hv + + use (χ₀.comp π₁) + c • (e.comp π₂) + refine nontrivial_of_ne ⟨v, ?_⟩ 0 ?_ + · rw [mem_weightSpace] + intro x + have hπ : (π₁ x : L) + π₂ x = x := linear_proj_add_linearProjOfIsCompl_eq_self hA x + suffices ⁅(π₂ x : L), v⁆ = (c • e (π₂ x)) • v by + calc ⁅x, v⁆ + = ⁅π₁ x, v⁆ + ⁅(π₂ x : L), v⁆ := congr(⁅$hπ.symm, v⁆) ▸ add_lie _ _ _ + _ = χ₀ (π₁ x) • v + (c • e (π₂ x)) • v := by rw [hv' (π₁ x), this] + _ = _ := by simp [add_smul] + calc ⁅(π₂ x : L), v⁆ + = e (π₂ x) • ↑(c • ⟨v, hv⟩ : W) := by rw [← he, smul_lie, ← hvc.apply_eq_smul]; rfl + _ = (c • e (π₂ x)) • v := by rw [smul_assoc, smul_comm]; rfl + · simpa [ne_eq, LieSubmodule.mk_eq_zero] using hvc.right + +variable (k L V) +variable [Nontrivial V] + +open LieAlgebra + +-- This lemma is the central inductive argument in the proof of Lie's theorem below. +-- The statement is identical to `LieModule.exists_forall_lie_eq_smul_of_isSolvable` +-- except that it additionally assumes a finiteness hypothesis. +private lemma exists_forall_lie_eq_smul_of_isSolvable_of_finite + (L : Type*) [LieRing L] [LieAlgebra k L] [LieRingModule L V] [LieModule k L V] + [IsSolvable k L] [LieModule.IsTriangularizable k L V] [Module.Finite k L] : + ∃ χ : Module.Dual k L, Nontrivial (weightSpace V χ) := by + obtain H|⟨A, hA, hAL⟩ := eq_top_or_exists_le_coatom (derivedSeries k L 1).toSubmodule + · obtain _|_ := subsingleton_or_nontrivial L + · use 0 + simpa [mem_weightSpace, nontrivial_iff] using exists_pair_ne V + · rw [LieSubmodule.toSubmodule_eq_top] at H + exact ((derivedSeries_lt_top_of_solvable k L).ne H).elim + lift A to LieIdeal k L + · intros + exact hAL <| LieSubmodule.lie_mem_lie (LieSubmodule.mem_top _) (LieSubmodule.mem_top _) + change LieIdeal k L at A -- remove this line when bug in `lift` is fixed (#15865) + obtain ⟨χ', _⟩ := exists_forall_lie_eq_smul_of_isSolvable_of_finite A + exact exists_nontrivial_weightSpace_of_lieIdeal A hA χ' +termination_by Module.finrank k L +decreasing_by + simp_wf + rw [← finrank_top k L] + apply Submodule.finrank_lt_finrank_of_lt + exact hA.lt_top + +/-- **Lie's theorem**: Lie modules of solvable Lie algebras over fields of characteristic 0 +have a common eigenvector for the action of all elements of the Lie algebra. + +See `LieModule.exists_nontrivial_weightSpace_of_isNilpotent` for the variant that +assumes that `L` is nilpotent and drops the condition that `k` is of characteristic zero. -/ +theorem exists_nontrivial_weightSpace_of_isSolvable + [IsSolvable k L] [LieModule.IsTriangularizable k L V] : + ∃ χ : Module.Dual k L, Nontrivial (weightSpace V χ) := by + let imL := (toEnd k L V).range + let toEndo : L →ₗ[k] imL := LinearMap.codRestrict imL.toSubmodule (toEnd k L V) + (fun x ↦ LinearMap.mem_range.mpr ⟨x, rfl⟩ : ∀ x : L, (toEnd k L V) x ∈ imL) + have ⟨χ, h⟩ := exists_forall_lie_eq_smul_of_isSolvable_of_finite k V imL + use χ.comp toEndo + obtain ⟨⟨v, hv⟩, hv0⟩ := exists_ne (0 : weightSpace V χ) + refine nontrivial_of_ne ⟨v, ?_⟩ 0 ?_ + · rw [mem_weightSpace] at hv ⊢ + intro x + apply hv (toEndo x) + · simpa using hv0 + +end + +end LieModule diff --git a/Mathlib/Algebra/Lie/Submodule.lean b/Mathlib/Algebra/Lie/Submodule.lean index 7b1a08a9b5d90..c7f9b4f034571 100644 --- a/Mathlib/Algebra/Lie/Submodule.lean +++ b/Mathlib/Algebra/Lie/Submodule.lean @@ -77,6 +77,10 @@ instance (priority := high) coeSort : CoeSort (LieSubmodule R L M) (Type w) wher instance (priority := mid) coeSubmodule : CoeOut (LieSubmodule R L M) (Submodule R M) := ⟨toSubmodule⟩ +instance : CanLift (Submodule R M) (LieSubmodule R L M) (·) + (fun N ↦ ∀ {x : L} {m : M}, m ∈ N → ⁅x, m⁆ ∈ N) where + prf N hN := ⟨⟨N, hN⟩, rfl⟩ + @[norm_cast] theorem coe_toSubmodule : ((N : Submodule R M) : Set M) = N := rfl diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean index 40e83c9dbe56d..e5914475629d3 100644 --- a/Mathlib/Algebra/Lie/TraceForm.lean +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -94,8 +94,7 @@ 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 - [Module.Free R M] [Module.Finite R M] [IsReduced R] [IsNilpotent R L M] : +@[simp] lemma traceForm_eq_zero_of_isNilpotent [IsReduced R] [IsNilpotent R L M] : traceForm R L M = 0 := by ext x y simp only [traceForm_apply_apply, LinearMap.zero_apply, ← isNilpotent_iff_eq_zero] diff --git a/Mathlib/Algebra/Lie/Weights/Linear.lean b/Mathlib/Algebra/Lie/Weights/Linear.lean index 596b044234646..bd48f9a203f78 100644 --- a/Mathlib/Algebra/Lie/Weights/Linear.lean +++ b/Mathlib/Algebra/Lie/Weights/Linear.lean @@ -41,7 +41,7 @@ or `R` has characteristic zero. open Set -variable (R L M : Type*) [CommRing R] [LieRing L] [LieAlgebra 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 @@ -244,4 +244,19 @@ lemma exists_forall_lie_eq_smul [LinearWeights R L M] [IsNoetherian R M] (χ : W have := hm₁ x rwa [coe_lie_shiftedGenWeightSpace_apply, sub_eq_zero] at this +/-- See `LieModule.exists_nontrivial_weightSpace_of_isSolvable` for the variant that +only assumes that `L` is solvable but additionally requires `k` to be of characteristic zero. -/ +lemma exists_nontrivial_weightSpace_of_isNilpotent [Field k] [LieAlgebra k L] [Module k M] + [Module.Finite k M] [LieModule k L M] [LieAlgebra.IsNilpotent k L] [LinearWeights k L M] + [IsTriangularizable k L M] [Nontrivial M] : + ∃ χ : Module.Dual k L, Nontrivial (weightSpace M χ) := by + obtain ⟨χ⟩ : Nonempty (Weight k L M) := by + by_contra contra + rw [not_nonempty_iff] at contra + simpa only [iSup_of_empty, bot_ne_top] using LieModule.iSup_genWeightSpace_eq_top' k L M + obtain ⟨m, hm₀, hm⟩ := exists_forall_lie_eq_smul k L M χ + simp only [LieSubmodule.nontrivial_iff_ne_bot, LieSubmodule.eq_bot_iff, Weight.coe_coe, ne_eq, + not_forall, Classical.not_imp] + exact ⟨χ.toLinear, m, by simpa [mem_weightSpace], hm₀⟩ + end LieModule diff --git a/Mathlib/LinearAlgebra/Trace.lean b/Mathlib/LinearAlgebra/Trace.lean index c64e1436ec160..08b144fd92f98 100644 --- a/Mathlib/LinearAlgebra/Trace.lean +++ b/Mathlib/LinearAlgebra/Trace.lean @@ -292,15 +292,18 @@ 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 + by_cases H : ∃ s : Finset M, Nonempty (Basis s R M) + swap + · rw [LinearMap.trace, dif_neg H] + exact IsNilpotent.zero + obtain ⟨s, ⟨b⟩⟩ := H + classical rw [trace_eq_matrix_trace R b] 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/RingTheory/Trace/Basic.lean b/Mathlib/RingTheory/Trace/Basic.lean index 9e7cad92167ba..dbe13c0d621bf 100644 --- a/Mathlib/RingTheory/Trace/Basic.lean +++ b/Mathlib/RingTheory/Trace/Basic.lean @@ -494,14 +494,8 @@ namespace Algebra /-- The trace of a nilpotent element is nilpotent. -/ lemma trace_isNilpotent_of_isNilpotent {R S : Type*} [CommRing R] [CommRing S] [Algebra R S] {x : S} - (hx : IsNilpotent x) : IsNilpotent (trace R S x) := by - by_cases hS : ∃ s : Finset S, Nonempty (Basis s R S) - · obtain ⟨s, ⟨b⟩⟩ := hS - have := Module.Finite.of_basis b - have := (Module.free_def R S).mpr ⟨s, ⟨b⟩⟩ - apply LinearMap.isNilpotent_trace_of_isNilpotent (hx.map (lmul R S)) - · rw [trace_eq_zero_of_not_exists_basis _ hS, LinearMap.zero_apply] - exact IsNilpotent.zero + (hx : IsNilpotent x) : IsNilpotent (trace R S x) := + LinearMap.isNilpotent_trace_of_isNilpotent (hx.map (lmul R S)) end Algebra From 2c7a747a2165e5b2a3e5a4490a24227cfd7fdb5e Mon Sep 17 00:00:00 2001 From: Seewoo Lee Date: Sat, 4 Jan 2025 18:40:14 +0000 Subject: [PATCH 142/189] feat(Polynomial): `(a^n)' = 0` iff `a' = 0` when `n` doesn't divide the characteristic (#20190) For a polynomial $a \in k[X]$ and $n$ with $n \ne 0$ in $k$, $(a^n)' = 0$ if and only if $a' = 0$. Co-authored-by: Seewoo Lee <49933279+seewoo5@users.noreply.github.com> --- Mathlib/Algebra/Polynomial/Derivative.lean | 13 ++++++++++++- 1 file changed, 12 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Polynomial/Derivative.lean b/Mathlib/Algebra/Polynomial/Derivative.lean index 9c219329bffc2..c23fdfc0c1e98 100644 --- a/Mathlib/Algebra/Polynomial/Derivative.lean +++ b/Mathlib/Algebra/Polynomial/Derivative.lean @@ -3,7 +3,6 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker -/ -import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Algebra.Polynomial.Degree.Domain import Mathlib.Algebra.Polynomial.Degree.Support import Mathlib.Algebra.Polynomial.Eval.Coeff @@ -661,6 +660,18 @@ theorem dvd_derivative_iff {P : R[X]} : P ∣ derivative P ↔ derivative P = 0 end NoZeroDivisors +section CommSemiringNoZeroDivisors + +variable [CommSemiring R] [NoZeroDivisors R] + +theorem derivative_pow_eq_zero {n : ℕ} (chn : (n : R) ≠ 0) {a : R[X]} : + derivative (a ^ n) = 0 ↔ derivative a = 0 := by + nontriviality R + rw [← C_ne_zero, C_eq_natCast] at chn + simp +contextual [derivative_pow, or_imp, chn] + +end CommSemiringNoZeroDivisors + end Derivative end Polynomial From a0b2b5768e28f3b54fc60d135e9497b8929e5fcc Mon Sep 17 00:00:00 2001 From: Seewoo Lee Date: Sat, 4 Jan 2025 18:40:15 +0000 Subject: [PATCH 143/189] =?UTF-8?q?feat(Radical):=20`(radical=20a).natDegr?= =?UTF-8?q?ee=20=E2=89=A4=20a.natDegree`=20(#20468)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/RingTheory/Polynomial/Radical.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/Mathlib/RingTheory/Polynomial/Radical.lean b/Mathlib/RingTheory/Polynomial/Radical.lean index 4a8f0a57af1e6..25b58667e4848 100644 --- a/Mathlib/RingTheory/Polynomial/Radical.lean +++ b/Mathlib/RingTheory/Polynomial/Radical.lean @@ -19,6 +19,15 @@ open Polynomial UniqueFactorizationMonoid UniqueFactorizationDomain EuclideanDom variable {k : Type*} [Field k] [DecidableEq k] +theorem degree_radical_le {a : k[X]} (h : a ≠ 0) : + (radical a).degree ≤ a.degree := degree_le_of_dvd (radical_dvd_self a) h + +theorem natDegree_radical_le {a : k[X]} : + (radical a).natDegree ≤ a.natDegree := by + by_cases ha : a = 0 + · simp [ha] + · exact natDegree_le_of_dvd (radical_dvd_self a) ha + theorem divRadical_dvd_derivative (a : k[X]) : divRadical a ∣ derivative a := by induction a using induction_on_coprime · case h0 => From b43248be90578db4b4158698029b1200ab4df641 Mon Sep 17 00:00:00 2001 From: Whysoserioushah Date: Sat, 4 Jan 2025 18:55:47 +0000 Subject: [PATCH 144/189] feat(Algebra): more on `OrthogonalIdempotents` (#18195) Co-authored-by : @jjaassoonn Co-authored-by: Whysoserioushah <109107491+Whysoserioushah@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Algebra/DirectSum/Idempotents.lean | 55 ++++++++++++++++++++++ Mathlib/RingTheory/Idempotents.lean | 1 - 3 files changed, 56 insertions(+), 1 deletion(-) create mode 100644 Mathlib/Algebra/DirectSum/Idempotents.lean diff --git a/Mathlib.lean b/Mathlib.lean index 9ccfe90d0fd2c..3a6ceab1f6a98 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -205,6 +205,7 @@ import Mathlib.Algebra.DirectSum.Algebra import Mathlib.Algebra.DirectSum.Basic import Mathlib.Algebra.DirectSum.Decomposition import Mathlib.Algebra.DirectSum.Finsupp +import Mathlib.Algebra.DirectSum.Idempotents import Mathlib.Algebra.DirectSum.Internal import Mathlib.Algebra.DirectSum.LinearMap import Mathlib.Algebra.DirectSum.Module diff --git a/Mathlib/Algebra/DirectSum/Idempotents.lean b/Mathlib/Algebra/DirectSum/Idempotents.lean new file mode 100644 index 0000000000000..bb76018d54b2c --- /dev/null +++ b/Mathlib/Algebra/DirectSum/Idempotents.lean @@ -0,0 +1,55 @@ +/- +Copyright (c) 2025 Yunzhou Xie. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yunzhou Xie, Jujian Zhang, Andrew Yang +-/ +import Mathlib.RingTheory.Idempotents +import Mathlib.Algebra.DirectSum.Decomposition + +/-! +# Decomposition of the identity of a ring into orthogonal idempotents + +In this file we show that if a ring `R` can be decomposed into a direct sum +of (left) ideals `R = V₁ ⊕ V₂ ⊕ ⬝ ⬝ ⬝ ⊕ Vₙ` then in the corresponding decomposition +`1 = e₁ + e₂ + ⬝ ⬝ ⬝ + eₙ` with `eᵢ ∈ Vᵢ`, each `eᵢ` is an idempotent and the +`eᵢ`'s form a family of complete orthogonal idempotents. +-/ + +namespace DirectSum + +section OrthogonalIdempotents + +variable {R I : Type*} [Ring R] [DecidableEq I] (V : I → Ideal R) [Decomposition V] + +/-- The decomposition of `(1 : R)` where `1 = e₁ + e₂ + ⬝ ⬝ ⬝ + eₙ` which is induced by + the decomposition of the ring `R = V1 ⊕ V2 ⊕ ⬝ ⬝ ⬝ ⊕ Vn`.-/ +def idempotent (i : I) : R := + decompose V 1 i + +lemma decompose_eq_mul_idempotent + (x : R) (i : I) : decompose V x i = x * idempotent V i := by + rw [← smul_eq_mul (a := x), idempotent, ← Submodule.coe_smul, ← smul_apply, ← decompose_smul, + smul_eq_mul, mul_one] + +lemma isIdempotentElem_idempotent (i : I) : IsIdempotentElem (idempotent V i : R) := by + rw [IsIdempotentElem, ← decompose_eq_mul_idempotent, idempotent, decompose_coe, of_eq_same] + +/-- If a ring can be decomposed into direct sum of finite left ideals `Vᵢ` + where `1 = e₁ + ... + eₙ` and `eᵢ ∈ Vᵢ`, then `eᵢ` is a family of complete + orthogonal idempotents.-/ +theorem completeOrthogonalIdempotents_idempotent [Fintype I]: + CompleteOrthogonalIdempotents fun i ↦ idempotent V i where + idem := isIdempotentElem_idempotent V + ortho i j hij := by + simp only + rw [← decompose_eq_mul_idempotent, idempotent, decompose_coe, + of_eq_of_ne (h := hij), Submodule.coe_zero] + complete := by + apply (decompose V).injective + refine DFunLike.ext _ _ fun i ↦ ?_ + rw [decompose_sum, DFinsupp.finset_sum_apply] + simp [idempotent, of_apply] + +end OrthogonalIdempotents + +end DirectSum diff --git a/Mathlib/RingTheory/Idempotents.lean b/Mathlib/RingTheory/Idempotents.lean index 65b8c4836a131..c7b091ec93c93 100644 --- a/Mathlib/RingTheory/Idempotents.lean +++ b/Mathlib/RingTheory/Idempotents.lean @@ -7,7 +7,6 @@ import Mathlib.Algebra.GeomSum import Mathlib.Algebra.Polynomial.AlgebraMap import Mathlib.RingTheory.Ideal.Quotient.Operations import Mathlib.RingTheory.Nilpotent.Defs -import Mathlib.Tactic.StacksAttribute /-! From 80fa5c2eaa8a55241925d327439d0d8085c68cc8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sat, 4 Jan 2025 21:33:19 +0000 Subject: [PATCH 145/189] chore(SetTheory/Ordinal/Arithmetic): `IsLeftCancelAdd Ordinal` (#19888) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In general, `AddLeftMono` + `AddLeftReflectLE` + `PartialOrder` → `IsLeftCancelAdd`, but this feels like too niche of a situation to create the general instance. --- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 17 ++++++++++------- Mathlib/SetTheory/Ordinal/FixedPoint.lean | 2 +- 2 files changed, 11 insertions(+), 8 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 7fd96b5c22941..ee879a0ea1dde 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -94,8 +94,12 @@ instance instAddLeftReflectLE : refine (RelEmbedding.ofMonotone g fun _ _ h ↦ ?_).ordinal_type_le rwa [← @Sum.lex_inr_inr _ t _ s, ← hg, ← hg, f.map_rel_iff, Sum.lex_inr_inr] -theorem add_left_cancel (a) {b c : Ordinal} : a + b = a + c ↔ b = c := by - simp only [le_antisymm_iff, add_le_add_iff_left] +instance : IsLeftCancelAdd Ordinal where + add_left_cancel a b c h := by simpa only [le_antisymm_iff, add_le_add_iff_left] using h + +@[deprecated add_left_cancel_iff (since := "2024-12-11")] +protected theorem add_left_cancel (a) {b c : Ordinal} : a + b = a + c ↔ b = c := + add_left_cancel_iff private theorem add_lt_add_iff_left' (a) {b c : Ordinal} : a + b < a + c ↔ b < c := by rw [← not_le, ← not_le, add_le_add_iff_left] @@ -2298,10 +2302,9 @@ alias nat_cast_pos := natCast_pos @[simp, norm_cast] theorem natCast_sub (m n : ℕ) : ((m - n : ℕ) : Ordinal) = m - n := by rcases le_total m n with h | h - · rw [tsub_eq_zero_iff_le.2 h, Ordinal.sub_eq_zero_iff_le.2 (Nat.cast_le.2 h)] - rfl - · apply (add_left_cancel n).1 - rw [← Nat.cast_add, add_tsub_cancel_of_le h, Ordinal.add_sub_cancel_of_le (Nat.cast_le.2 h)] + · rw [tsub_eq_zero_iff_le.2 h, Ordinal.sub_eq_zero_iff_le.2 (Nat.cast_le.2 h), Nat.cast_zero] + · rw [← add_left_cancel_iff (a := ↑n), ← Nat.cast_add, add_tsub_cancel_of_le h, + Ordinal.add_sub_cancel_of_le (Nat.cast_le.2 h)] @[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_sub := natCast_sub @@ -2323,7 +2326,7 @@ alias nat_cast_div := natCast_div @[simp, norm_cast] theorem natCast_mod (m n : ℕ) : ((m % n : ℕ) : Ordinal) = m % n := by - rw [← add_left_cancel, div_add_mod, ← natCast_div, ← natCast_mul, ← Nat.cast_add, + rw [← add_left_cancel_iff, div_add_mod, ← natCast_div, ← natCast_mul, ← Nat.cast_add, Nat.div_add_mod] @[deprecated "No deprecation message was provided." (since := "2024-04-17")] diff --git a/Mathlib/SetTheory/Ordinal/FixedPoint.lean b/Mathlib/SetTheory/Ordinal/FixedPoint.lean index 59dc4a1c4d06d..536e6e2466e85 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPoint.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPoint.lean @@ -688,7 +688,7 @@ theorem mul_eq_right_iff_opow_omega0_dvd {a b : Ordinal} : a * b = b ↔ a ^ ω refine ⟨fun hab => ?_, fun h => ?_⟩ · rw [dvd_iff_mod_eq_zero] rw [← div_add_mod b (a ^ ω), mul_add, ← mul_assoc, ← opow_one_add, one_add_omega0, - add_left_cancel] at hab + add_left_cancel_iff] at hab cases' eq_zero_or_opow_omega0_le_of_mul_eq_right hab with hab hab · exact hab refine (not_lt_of_le hab (mod_lt b (opow_ne_zero ω ?_))).elim From 1e25efce54c09ee2321f801cb4b51126869ae8c6 Mon Sep 17 00:00:00 2001 From: smorel394 Date: Sat, 4 Jan 2025 21:54:34 +0000 Subject: [PATCH 146/189] chore(Algebra/Category/MonCat/Colimits): remove smallness condition (#20473) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The file constructed colimits of functors `F : J ⥤ MonCat.{v}` under the assumption that `J` is a `SmallCategory.{v}`. It turns out that the smallness conditions on the `Hom` types of `J` is unnecessary and can simply be removed. --- Mathlib/Algebra/Category/MonCat/Colimits.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Category/MonCat/Colimits.lean b/Mathlib/Algebra/Category/MonCat/Colimits.lean index ac67635fefbbe..9e317a99bb362 100644 --- a/Mathlib/Algebra/Category/MonCat/Colimits.lean +++ b/Mathlib/Algebra/Category/MonCat/Colimits.lean @@ -47,7 +47,7 @@ Monoid.mk : {M : Type u} → -/ -universe v +universe v u open CategoryTheory @@ -63,7 +63,7 @@ and the identifications given by the morphisms in the diagram. -/ -variable {J : Type v} [SmallCategory J] (F : J ⥤ MonCat.{v}) +variable {J : Type v} [Category.{u} J] (F : J ⥤ MonCat.{v}) /-- An inductive type representing all monoid expressions (without relations) on a collection of types indexed by the objects of `J`. From d2a3986af5cd87ce2465108c4ecc7f074f1e88c9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Miyahara=20K=C5=8D?= Date: Sun, 5 Jan 2025 01:09:11 +0000 Subject: [PATCH 147/189] chore: remove unnecessary adaptation notes (#20467) This PR removes following unnecesarry adaptation notes. `simp_rw` to `rw` or `simp` * `Mathlib.Algebra.Group.Support` * `Mathlib.Analysis.Normed.Group.Basic` a local instance is not required anymore due to `set_option maxSynthPendingDepth 2` * `Mathlib.Analysis.NormedSpace.Multilinear.Basic` * `Mathlib.Analysis.NormedSpace.OperatorNorm.Bilinear` * `Mathlib.Analysis.NormedSpace.OperatorNorm.Mul` * `Mathlib.Analysis.NormedSpace.OperatorNorm.NormedSpace` --- Mathlib/Algebra/Group/Support.lean | 2 -- Mathlib/Analysis/Normed/Group/Basic.lean | 3 -- .../NormedSpace/Multilinear/Basic.lean | 10 ------ .../NormedSpace/OperatorNorm/Bilinear.lean | 34 ------------------- .../NormedSpace/OperatorNorm/Mul.lean | 9 ----- .../NormedSpace/OperatorNorm/NormedSpace.lean | 8 ----- 6 files changed, 66 deletions(-) diff --git a/Mathlib/Algebra/Group/Support.lean b/Mathlib/Algebra/Group/Support.lean index 6a3a98f9e54df..5af777a1bf618 100644 --- a/Mathlib/Algebra/Group/Support.lean +++ b/Mathlib/Algebra/Group/Support.lean @@ -110,8 +110,6 @@ theorem disjoint_mulSupport_iff {f : α → M} {s : Set α} : @[to_additive (attr := simp)] theorem mulSupport_eq_empty_iff {f : α → M} : mulSupport f = ∅ ↔ f = 1 := by - #adaptation_note /-- This used to be `simp_rw` rather than `rw`, - but this broke `to_additive` as of `nightly-2024-03-07` -/ rw [← subset_empty_iff, mulSupport_subset_iff', funext_iff] simp diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index ba1878a99efcb..fc2858dbb2006 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -957,9 +957,6 @@ theorem SeminormedCommGroup.mem_closure_iff : @[to_additive] theorem SeminormedGroup.tendstoUniformlyOn_one {f : ι → κ → G} {s : Set κ} {l : Filter ι} : TendstoUniformlyOn f 1 l s ↔ ∀ ε > 0, ∀ᶠ i in l, ∀ x ∈ s, ‖f i x‖ < ε := by - #adaptation_note /-- nightly-2024-03-11. - Originally this was `simp_rw` instead of `simp only`, - but this creates a bad proof term with nested `OfNat.ofNat` that trips up `@[to_additive]`. -/ simp only [tendstoUniformlyOn_iff, Pi.one_apply, dist_one_left] @[to_additive] diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 96a773a054737..6c60138a2473d 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -867,16 +867,6 @@ def smulRightL : ContinuousMultilinearMap 𝕜 E 𝕜 →L[𝕜] G →L[𝕜] Co @[simp] lemma smulRightL_apply (f : ContinuousMultilinearMap 𝕜 E 𝕜) (z : G) : smulRightL 𝕜 E G f z = f.smulRight z := rfl -#adaptation_note -/-- -Before https://github.com/leanprover/lean4/pull/4119 we had to create a local instance: -``` -letI : SeminormedAddCommGroup - (ContinuousMultilinearMap 𝕜 E 𝕜 →L[𝕜] G →L[𝕜] ContinuousMultilinearMap 𝕜 E G) := - ContinuousLinearMap.toSeminormedAddCommGroup - (F := G →L[𝕜] ContinuousMultilinearMap 𝕜 E G) (σ₁₂ := RingHom.id 𝕜) -``` --/ set_option maxSynthPendingDepth 2 in lemma norm_smulRightL_le : ‖smulRightL 𝕜 E G‖ ≤ 1 := LinearMap.mkContinuous₂_norm_le _ zero_le_one _ diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean index 6b959b42931e5..9cb3ac62feea8 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean @@ -273,14 +273,6 @@ def compSL : (F →SL[σ₂₃] G) →L[𝕜₃] (E →SL[σ₁₂] F) →SL[σ Pi.smul_apply]) 1 fun f g => by simpa only [one_mul] using opNorm_comp_le f g -#adaptation_note -/-- -Before https://github.com/leanprover/lean4/pull/4119 we had to create a local instance: -``` -letI : Norm ((F →SL[σ₂₃] G) →L[𝕜₃] (E →SL[σ₁₂] F) →SL[σ₂₃] E →SL[σ₁₃] G) := - hasOpNorm (𝕜₂ := 𝕜₃) (E := F →SL[σ₂₃] G) (F := (E →SL[σ₁₂] F) →SL[σ₂₃] E →SL[σ₁₃] G) -``` --/ set_option maxSynthPendingDepth 2 in theorem norm_compSL_le : ‖compSL E F G σ₁₂ σ₂₃‖ ≤ 1 := LinearMap.mkContinuous₂_norm_le _ zero_le_one _ @@ -309,14 +301,6 @@ variable (𝕜 σ₁₂ σ₂₃ E Fₗ Gₗ) def compL : (Fₗ →L[𝕜] Gₗ) →L[𝕜] (E →L[𝕜] Fₗ) →L[𝕜] E →L[𝕜] Gₗ := compSL E Fₗ Gₗ (RingHom.id 𝕜) (RingHom.id 𝕜) -#adaptation_note -/-- -Before https://github.com/leanprover/lean4/pull/4119 we had to create a local instance: -``` -letI : Norm ((Fₗ →L[𝕜] Gₗ) →L[𝕜] (E →L[𝕜] Fₗ) →L[𝕜] E →L[𝕜] Gₗ) := - hasOpNorm (𝕜₂ := 𝕜) (E := Fₗ →L[𝕜] Gₗ) (F := (E →L[𝕜] Fₗ) →L[𝕜] E →L[𝕜] Gₗ) -``` --/ set_option maxSynthPendingDepth 2 in theorem norm_compL_le : ‖compL 𝕜 E Fₗ Gₗ‖ ≤ 1 := norm_compSL_le _ _ _ _ _ @@ -339,15 +323,6 @@ def precompL (L : E →L[𝕜] Fₗ →L[𝕜] Gₗ) : (Eₗ →L[𝕜] E) →L[ @[simp] lemma precompL_apply (L : E →L[𝕜] Fₗ →L[𝕜] Gₗ) (u : Eₗ →L[𝕜] E) (f : Fₗ) (g : Eₗ) : precompL Eₗ L u f g = L (u g) f := rfl -#adaptation_note -/-- -Before https://github.com/leanprover/lean4/pull/4119 -we had to create a local instance in the signature: -``` -letI : SeminormedAddCommGroup ((Eₗ →L[𝕜] Fₗ) →L[𝕜] Eₗ →L[𝕜] Gₗ) := inferInstance -letI : NormedSpace 𝕜 ((Eₗ →L[𝕜] Fₗ) →L[𝕜] Eₗ →L[𝕜] Gₗ) := inferInstance -``` --/ set_option maxSynthPendingDepth 2 in theorem norm_precompR_le (L : E →L[𝕜] Fₗ →L[𝕜] Gₗ) : ‖precompR Eₗ L‖ ≤ ‖L‖ := calc @@ -355,15 +330,6 @@ theorem norm_precompR_le (L : E →L[𝕜] Fₗ →L[𝕜] Gₗ) : ‖precompR E _ ≤ 1 * ‖L‖ := mul_le_mul_of_nonneg_right (norm_compL_le _ _ _ _) (norm_nonneg L) _ = ‖L‖ := by rw [one_mul] -#adaptation_note -/-- -Before https://github.com/leanprover/lean4/pull/4119 -we had to create a local instance in the signature: -``` -letI : Norm ((Eₗ →L[𝕜] E) →L[𝕜] Fₗ →L[𝕜] Eₗ →L[𝕜] Gₗ) := - hasOpNorm (𝕜₂ := 𝕜) (E := Eₗ →L[𝕜] E) (F := Fₗ →L[𝕜] Eₗ →L[𝕜] Gₗ) -``` --/ set_option maxSynthPendingDepth 2 in theorem norm_precompL_le (L : E →L[𝕜] Fₗ →L[𝕜] Gₗ) : ‖precompL Eₗ L‖ ≤ ‖L‖ := by rw [precompL, opNorm_flip, ← opNorm_flip L] diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean index 25796540eece9..c45a659de2c23 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean @@ -97,15 +97,6 @@ theorem opNorm_mulLeftRight_apply_le (x : 𝕜') : ‖mulLeftRight 𝕜 𝕜' x @[deprecated (since := "2024-02-02")] alias op_norm_mulLeftRight_apply_le := opNorm_mulLeftRight_apply_le -#adaptation_note -/-- -Before https://github.com/leanprover/lean4/pull/4119 -we had to create a local instance in the signature: -``` -letI : Norm (𝕜' →L[𝕜] 𝕜' →L[𝕜] 𝕜' →L[𝕜] 𝕜') := - hasOpNorm (𝕜₂ := 𝕜) (E := 𝕜') (F := 𝕜' →L[𝕜] 𝕜' →L[𝕜] 𝕜') -``` --/ set_option maxSynthPendingDepth 2 in theorem opNorm_mulLeftRight_le : ‖mulLeftRight 𝕜 𝕜'‖ ≤ 1 := diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean index b1ab15b72ee5d..b08ba685f3320 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean @@ -215,14 +215,6 @@ alias op_norm_comp_linearIsometryEquiv := opNorm_comp_linearIsometryEquiv theorem norm_smulRightL (c : E →L[𝕜] 𝕜) [Nontrivial Fₗ] : ‖smulRightL 𝕜 E Fₗ c‖ = ‖c‖ := ContinuousLinearMap.homothety_norm _ c.norm_smulRight_apply -#adaptation_note -/-- -Before https://github.com/leanprover/lean4/pull/4119 we had to create a local instance: -``` -letI : SeminormedAddCommGroup ((E →L[𝕜] 𝕜) →L[𝕜] Fₗ →L[𝕜] E →L[𝕜] Fₗ) := - toSeminormedAddCommGroup (F := Fₗ →L[𝕜] E →L[𝕜] Fₗ) (𝕜 := 𝕜) (σ₁₂ := RingHom.id 𝕜) -``` --/ set_option maxSynthPendingDepth 2 in lemma norm_smulRightL_le : ‖smulRightL 𝕜 E Fₗ‖ ≤ 1 := LinearMap.mkContinuous₂_norm_le _ zero_le_one _ From a5ae6711d1a1d4b81f729eb701161c24db3a6d9d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 5 Jan 2025 01:44:22 +0000 Subject: [PATCH 148/189] feat: shorthand lemmas for the L1 norm (#20383) These are consistent with the existing lemmas about the L2 norm. --- Mathlib/Analysis/Matrix.lean | 2 +- Mathlib/Analysis/Normed/Lp/PiLp.lean | 39 ++++++++++++++++++++------ Mathlib/Analysis/Normed/Lp/ProdLp.lean | 35 +++++++++++++++++++++-- 3 files changed, 65 insertions(+), 11 deletions(-) diff --git a/Mathlib/Analysis/Matrix.lean b/Mathlib/Analysis/Matrix.lean index e691c6be69c1d..7a1f22fb8be30 100644 --- a/Mathlib/Analysis/Matrix.lean +++ b/Mathlib/Analysis/Matrix.lean @@ -246,7 +246,7 @@ theorem linfty_opNorm_def (A : Matrix m n α) : ‖A‖ = ((Finset.univ : Finset m).sup fun i : m => ∑ j : n, ‖A i j‖₊ : ℝ≥0) := by -- Porting note: added change ‖fun i => (WithLp.equiv 1 _).symm (A i)‖ = _ - simp [Pi.norm_def, PiLp.nnnorm_eq_sum ENNReal.one_ne_top] + simp [Pi.norm_def, PiLp.nnnorm_eq_of_L1] @[deprecated (since := "2024-02-02")] alias linfty_op_norm_def := linfty_opNorm_def diff --git a/Mathlib/Analysis/Normed/Lp/PiLp.lean b/Mathlib/Analysis/Normed/Lp/PiLp.lean index 3b9cc58963b91..0c311a281bb35 100644 --- a/Mathlib/Analysis/Normed/Lp/PiLp.lean +++ b/Mathlib/Analysis/Normed/Lp/PiLp.lean @@ -575,15 +575,37 @@ theorem norm_eq_of_nat {p : ℝ≥0∞} [Fact (1 ≤ p)] {β : ι → Type*} simp only [one_div, h, Real.rpow_natCast, ENNReal.toReal_nat, eq_self_iff_true, Finset.sum_congr, norm_eq_sum this] -theorem norm_eq_of_L2 {β : ι → Type*} [∀ i, SeminormedAddCommGroup (β i)] (x : PiLp 2 β) : +section L1 +variable {β} [∀ i, SeminormedAddCommGroup (β i)] + +theorem norm_eq_of_L1 (x : PiLp 1 β) : ‖x‖ = ∑ i : ι, ‖x i‖ := by + simp [norm_eq_sum] + +theorem nnnorm_eq_of_L1 (x : PiLp 1 β) : ‖x‖₊ = ∑ i : ι, ‖x i‖₊ := + NNReal.eq <| by push_cast; exact norm_eq_of_L1 x + +theorem dist_eq_of_L1 (x y : PiLp 1 β) : dist x y = ∑ i, dist (x i) (y i) := by + simp_rw [dist_eq_norm, norm_eq_of_L1, sub_apply] + +theorem nndist_eq_of_L1 (x y : PiLp 1 β) : nndist x y = ∑ i, nndist (x i) (y i) := + NNReal.eq <| by push_cast; exact dist_eq_of_L1 _ _ + +theorem edist_eq_of_L1 (x y : PiLp 1 β) : edist x y = ∑ i, edist (x i) (y i) := by + simp [PiLp.edist_eq_sum] + +end L1 + +section L2 +variable {β} [∀ i, SeminormedAddCommGroup (β i)] + +theorem norm_eq_of_L2 (x : PiLp 2 β) : ‖x‖ = √(∑ i : ι, ‖x i‖ ^ 2) := by - rw [norm_eq_of_nat 2 (by norm_cast) _] -- Porting note: was `convert` + rw [norm_eq_of_nat 2 (by norm_cast) _] rw [Real.sqrt_eq_rpow] norm_cast -theorem nnnorm_eq_of_L2 {β : ι → Type*} [∀ i, SeminormedAddCommGroup (β i)] (x : PiLp 2 β) : +theorem nnnorm_eq_of_L2 (x : PiLp 2 β) : ‖x‖₊ = NNReal.sqrt (∑ i : ι, ‖x i‖₊ ^ 2) := - -- Porting note: was `Subtype.ext` NNReal.eq <| by push_cast exact norm_eq_of_L2 x @@ -594,20 +616,21 @@ theorem norm_sq_eq_of_L2 (β : ι → Type*) [∀ i, SeminormedAddCommGroup (β simpa only [NNReal.coe_sum] using congr_arg ((↑) : ℝ≥0 → ℝ) this rw [nnnorm_eq_of_L2, NNReal.sq_sqrt] -theorem dist_eq_of_L2 {β : ι → Type*} [∀ i, SeminormedAddCommGroup (β i)] (x y : PiLp 2 β) : +theorem dist_eq_of_L2 (x y : PiLp 2 β) : dist x y = √(∑ i, dist (x i) (y i) ^ 2) := by simp_rw [dist_eq_norm, norm_eq_of_L2, sub_apply] -theorem nndist_eq_of_L2 {β : ι → Type*} [∀ i, SeminormedAddCommGroup (β i)] (x y : PiLp 2 β) : +theorem nndist_eq_of_L2 (x y : PiLp 2 β) : nndist x y = NNReal.sqrt (∑ i, nndist (x i) (y i) ^ 2) := - -- Porting note: was `Subtype.ext` NNReal.eq <| by push_cast exact dist_eq_of_L2 _ _ -theorem edist_eq_of_L2 {β : ι → Type*} [∀ i, SeminormedAddCommGroup (β i)] (x y : PiLp 2 β) : +theorem edist_eq_of_L2 (x y : PiLp 2 β) : edist x y = (∑ i, edist (x i) (y i) ^ 2) ^ (1 / 2 : ℝ) := by simp [PiLp.edist_eq_sum] +end L2 + instance instBoundedSMul [SeminormedRing 𝕜] [∀ i, SeminormedAddCommGroup (β i)] [∀ i, Module 𝕜 (β i)] [∀ i, BoundedSMul 𝕜 (β i)] : BoundedSMul 𝕜 (PiLp p β) := diff --git a/Mathlib/Analysis/Normed/Lp/ProdLp.lean b/Mathlib/Analysis/Normed/Lp/ProdLp.lean index 566fa89b049f3..d94faf96a1d7e 100644 --- a/Mathlib/Analysis/Normed/Lp/ProdLp.lean +++ b/Mathlib/Analysis/Normed/Lp/ProdLp.lean @@ -627,6 +627,36 @@ theorem prod_nnnorm_eq_sup (f : WithLp ∞ (α × β)) : ‖f‖₊ = ‖f.fst @[simp] theorem prod_norm_equiv_symm (f : α × β) : ‖(WithLp.equiv ⊤ _).symm f‖ = ‖f‖ := (prod_norm_equiv _).symm +section L1 + +theorem prod_norm_eq_of_L1 (x : WithLp 1 (α × β)) : + ‖x‖ = ‖x.fst‖ + ‖x.snd‖ := by + simp [prod_norm_eq_add] + +theorem prod_nnnorm_eq_of_L1 (x : WithLp 1 (α × β)) : + ‖x‖₊ = ‖x.fst‖₊ + ‖x.snd‖₊ := + NNReal.eq <| by + push_cast + exact prod_norm_eq_of_L1 x + +theorem prod_dist_eq_of_L1 (x y : WithLp 1 (α × β)) : + dist x y = dist x.fst y.fst + dist x.snd y.snd := by + simp_rw [dist_eq_norm, prod_norm_eq_of_L1, sub_fst, sub_snd] + +theorem prod_nndist_eq_of_L1 (x y : WithLp 1 (α × β)) : + nndist x y = nndist x.fst y.fst + nndist x.snd y.snd := + NNReal.eq <| by + push_cast + exact prod_dist_eq_of_L1 _ _ + +theorem prod_edist_eq_of_L1 (x y : WithLp 1 (α × β)) : + edist x y = edist x.fst y.fst + edist x.snd y.snd := by + simp [prod_edist_eq_add] + +end L1 + +section L2 + theorem prod_norm_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ = √(‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2) := by rw [prod_norm_eq_of_nat 2 (by norm_cast) _, Real.sqrt_eq_rpow] @@ -645,8 +675,7 @@ theorem prod_norm_sq_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ ^ 2 = ‖x.fst theorem prod_dist_eq_of_L2 (x y : WithLp 2 (α × β)) : dist x y = √(dist x.fst y.fst ^ 2 + dist x.snd y.snd ^ 2) := by - simp_rw [dist_eq_norm, prod_norm_eq_of_L2] - rfl + simp_rw [dist_eq_norm, prod_norm_eq_of_L2, sub_fst, sub_snd] theorem prod_nndist_eq_of_L2 (x y : WithLp 2 (α × β)) : nndist x y = NNReal.sqrt (nndist x.fst y.fst ^ 2 + nndist x.snd y.snd ^ 2) := @@ -658,6 +687,8 @@ theorem prod_edist_eq_of_L2 (x y : WithLp 2 (α × β)) : edist x y = (edist x.fst y.fst ^ 2 + edist x.snd y.snd ^ 2) ^ (1 / 2 : ℝ) := by simp [prod_edist_eq_add] +end L2 + end norm_of variable [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] From 9837ca9d65d9de6fad1ef4381750ca688774e608 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 5 Jan 2025 01:44:23 +0000 Subject: [PATCH 149/189] chore: bot validates lean-toolchain on bump/v4.X.0 branches (#20478) --- .github/workflows/nightly_detect_failure.yml | 43 ++++++++++++++++---- 1 file changed, 34 insertions(+), 9 deletions(-) diff --git a/.github/workflows/nightly_detect_failure.yml b/.github/workflows/nightly_detect_failure.yml index b37d8c9217e53..99a509a7466c9 100644 --- a/.github/workflows/nightly_detect_failure.yml +++ b/.github/workflows/nightly_detect_failure.yml @@ -170,15 +170,40 @@ jobs: uses: actions/github-script@v7 with: script: | - const response = await github.rest.repos.getContent({ - owner: context.repo.owner, - repo: context.repo.repo, - path: 'lean-toolchain', - ref: ${{ steps.latest_bump_branch.outputs.result }} - }); - const content = Buffer.from(response.data.content, 'base64').toString(); - const version = content.match(/leanprover\/lean4:nightly-(\d{4}-\d{2}-\d{2})/)[1]; - return version; + try { + const response = await github.rest.repos.getContent({ + owner: context.repo.owner, + repo: context.repo.repo, + path: 'lean-toolchain', + ref: ${{ steps.latest_bump_branch.outputs.result }} + }); + const content = Buffer.from(response.data.content, 'base64').toString(); + const match = content.match(/leanprover\/lean4:nightly-(\d{4}-\d{2}-\d{2})/); + if (!match) { + core.setFailed('Toolchain pattern did not match'); + core.setOutput('toolchain_content', content); + return null; + } + return match[1]; + } catch (error) { + core.setFailed(error.message); + return null; + } + + - name: Send warning message on Zulip if pattern doesn't match + if: failure() + uses: zulip/github-actions-zulip/send-message@v1 + with: + api-key: ${{ secrets.ZULIP_API_KEY }} + email: 'github-mathlib4-bot@leanprover.zulipchat.com' + organization-url: 'https://leanprover.zulipchat.com' + to: 'nightly-testing' + type: 'stream' + topic: 'Mathlib status updates' + content: | + ⚠️ Warning: The lean-toolchain file in the latest bump branch does not match the expected pattern 'leanprover/lean4:nightly-YYYY-MM-DD'. + Current content: ${{ steps.bump_version.outputs.toolchain_content }} + This needs to be fixed for the nightly testing process to work correctly. - name: Compare versions and post a reminder. env: From 0f1810248782d0843a813d22979f4fa77f0f9db5 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 5 Jan 2025 03:32:19 +0000 Subject: [PATCH 150/189] chore: bump toolchain to v4.16.0-rc1, and merge bump/v4.16.0 (#20464) Co-authored-by: Johan Commelin Co-authored-by: leanprover-community-mathlib4-bot --- Archive/Imo/Imo2024Q5.lean | 2 +- Mathlib.lean | 1 - Mathlib/Algebra/BigOperators/Fin.lean | 2 +- Mathlib/Algebra/BigOperators/Group/List.lean | 5 - .../Algebra/MvPolynomial/SchwartzZippel.lean | 2 +- .../EllipticCurve/NormalForms.lean | 3 +- .../Calculus/ContDiff/FaaDiBruno.lean | 8 +- Mathlib/Analysis/Convex/Continuous.lean | 3 +- .../CategoryTheory/Galois/Decomposition.lean | 10 +- Mathlib/Computability/TMToPartrec.lean | 37 ++- Mathlib/Computability/TuringMachine.lean | 11 +- Mathlib/Data/Fin/Basic.lean | 20 +- Mathlib/Data/Fin/Tuple/Basic.lean | 14 +- Mathlib/Data/Fin/Tuple/Take.lean | 2 +- Mathlib/Data/Int/Bitwise.lean | 2 - Mathlib/Data/Int/Lemmas.lean | 3 - Mathlib/Data/List/Basic.lean | 9 +- Mathlib/Data/List/Chain.lean | 4 +- Mathlib/Data/List/Lemmas.lean | 14 +- Mathlib/Data/List/Lex.lean | 37 +-- Mathlib/Data/List/MinMax.lean | 2 +- Mathlib/Data/List/Permutation.lean | 8 +- Mathlib/Data/Nat/Bitwise.lean | 5 +- Mathlib/Data/Nat/Defs.lean | 5 +- Mathlib/Data/Num/Basic.lean | 2 +- .../QPF/Multivariate/Constructions/Fix.lean | 2 +- Mathlib/Data/String/Basic.lean | 4 +- Mathlib/Data/Vector/Basic.lean | 5 +- Mathlib/Deprecated/LazyList.lean | 2 +- Mathlib/FieldTheory/Finite/GaloisField.lean | 4 +- Mathlib/Lean/Expr/Basic.lean | 2 +- .../Lean/Meta/RefinedDiscrTree/Encode.lean | 2 +- Mathlib/LinearAlgebra/Dual.lean | 2 +- Mathlib/LinearAlgebra/LinearIndependent.lean | 4 +- .../LinearAlgebra/Matrix/SemiringInverse.lean | 2 +- Mathlib/Logic/Basic.lean | 2 +- .../Integral/FundThmCalculus.lean | 2 +- Mathlib/MeasureTheory/Integral/Marginal.lean | 2 +- Mathlib/ModelTheory/Semantics.lean | 6 +- Mathlib/NumberTheory/Bernoulli.lean | 2 +- Mathlib/NumberTheory/Modular.lean | 5 +- Mathlib/Order/Fin/Basic.lean | 4 +- Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean | 3 +- Mathlib/RingTheory/Kaehler/JacobiZariski.lean | 2 +- Mathlib/RingTheory/Polynomial/Chebyshev.lean | 2 +- Mathlib/SetTheory/Ordinal/Notation.lean | 48 ++-- Mathlib/Tactic.lean | 1 - Mathlib/Tactic/CC/Addition.lean | 12 +- Mathlib/Tactic/CC/Datatypes.lean | 8 +- Mathlib/Tactic/CC/MkProof.lean | 12 +- Mathlib/Tactic/Common.lean | 2 - Mathlib/Tactic/DeriveToExpr.lean | 230 ------------------ Mathlib/Tactic/ITauto.lean | 1 - Mathlib/Tactic/Linarith/Frontend.lean | 4 +- Mathlib/Tactic/Linter/FlexibleLinter.lean | 2 +- Mathlib/Tactic/Linter/HaveLetLinter.lean | 3 +- Mathlib/Tactic/SuccessIfFailWithMsg.lean | 2 +- Mathlib/Tactic/ToExpr.lean | 45 +--- Mathlib/Tactic/ToLevel.lean | 25 -- .../Topology/Category/Profinite/Nobeling.lean | 7 +- MathlibTest/DeriveToExpr.lean | 3 +- MathlibTest/HashCommandLinter.lean | 6 + MathlibTest/HaveLetLinter.lean | 6 + MathlibTest/TransImports.lean | 4 +- MathlibTest/toAdditive.lean | 2 +- lake-manifest.json | 20 +- lakefile.lean | 10 +- lean-toolchain | 2 +- scripts/noshake.json | 4 +- 69 files changed, 187 insertions(+), 542 deletions(-) delete mode 100644 Mathlib/Tactic/DeriveToExpr.lean diff --git a/Archive/Imo/Imo2024Q5.lean b/Archive/Imo/Imo2024Q5.lean index 4fa8394b08d3d..04b40fc9ee4e2 100644 --- a/Archive/Imo/Imo2024Q5.lean +++ b/Archive/Imo/Imo2024Q5.lean @@ -302,7 +302,7 @@ lemma Path.tail_induction {motive : Path N → Prop} (ind : ∀ p, motive p.tail case cons head tail hi => by_cases h : (p'.cells[1]'p'.one_lt_length_cells).1 = 0 · refine ind p' ?_ - simp_rw [Path.tail, if_pos h, List.tail_cons] + simp_rw [Path.tail, if_pos h, p', List.tail_cons] exact hi _ _ _ _ · exact base p' h diff --git a/Mathlib.lean b/Mathlib.lean index 3a6ceab1f6a98..f692eadfdec90 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4841,7 +4841,6 @@ import Mathlib.Tactic.DefEqTransformations import Mathlib.Tactic.DeprecateTo import Mathlib.Tactic.DeriveCountable import Mathlib.Tactic.DeriveFintype -import Mathlib.Tactic.DeriveToExpr import Mathlib.Tactic.DeriveTraversable import Mathlib.Tactic.Eqns import Mathlib.Tactic.Eval diff --git a/Mathlib/Algebra/BigOperators/Fin.lean b/Mathlib/Algebra/BigOperators/Fin.lean index f336a553c7f57..045c79a9465a3 100644 --- a/Mathlib/Algebra/BigOperators/Fin.lean +++ b/Mathlib/Algebra/BigOperators/Fin.lean @@ -162,7 +162,7 @@ theorem prod_Ioi_succ {M : Type*} [CommMonoid M] {n : ℕ} (i : Fin n) (v : Fin @[to_additive] theorem prod_congr' {M : Type*} [CommMonoid M] {a b : ℕ} (f : Fin b → M) (h : a = b) : - (∏ i : Fin a, f (cast h i)) = ∏ i : Fin b, f i := by + (∏ i : Fin a, f (i.cast h)) = ∏ i : Fin b, f i := by subst h congr diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 1dd09cee52853..5ff4cb25da33e 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -5,7 +5,6 @@ Authors: Johannes Hölzl, Floris van Doorn, Sébastien Gouëzel, Alex J. Best -/ import Mathlib.Algebra.Divisibility.Basic import Mathlib.Algebra.Group.Int -import Mathlib.Data.List.Lemmas import Mathlib.Data.List.Dedup import Mathlib.Data.List.Flatten import Mathlib.Data.List.Pairwise @@ -637,10 +636,6 @@ end MonoidHom end MonoidHom -set_option linter.deprecated false in -@[simp, deprecated "No deprecation message was provided." (since := "2024-10-17")] -lemma Nat.sum_eq_listSum (l : List ℕ) : Nat.sum l = l.sum := rfl - namespace List lemma length_sigma {σ : α → Type*} (l₁ : List α) (l₂ : ∀ a, List (σ a)) : diff --git a/Mathlib/Algebra/MvPolynomial/SchwartzZippel.lean b/Mathlib/Algebra/MvPolynomial/SchwartzZippel.lean index 9b5e8cdc2e08b..05caccfa0a1a7 100644 --- a/Mathlib/Algebra/MvPolynomial/SchwartzZippel.lean +++ b/Mathlib/Algebra/MvPolynomial/SchwartzZippel.lean @@ -146,7 +146,7 @@ lemma schwartz_zippel_sup_sum : calc #{x₀ ∈ S 0 | eval (cons x₀ xₜ) p = 0} ≤ #pₓ.roots.toFinset := by gcongr - simp (config := { contextual := true }) [subset_iff, eval_eq_eval_mv_eval', pₓ, hpₓ₀] + simp +contextual [subset_iff, eval_eq_eval_mv_eval', pₓ, hpₓ₀, p'] _ ≤ Multiset.card pₓ.roots := pₓ.roots.toFinset_card_le _ ≤ pₓ.natDegree := pₓ.card_roots' _ = k := hpₓdeg diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/NormalForms.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/NormalForms.lean index 5738f216555e7..0bce274b8d038 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/NormalForms.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/NormalForms.lean @@ -425,7 +425,8 @@ theorem toCharThreeNF_spec_of_b₂_ne_zero (hb₂ : W.b₂ ≠ 0) : constructor · simp · simp - · field_simp [W.toShortNFOfCharThree_a₂ ▸ hb₂] + · have ha₂ : W'.a₂ ≠ 0 := W.toShortNFOfCharThree_a₂ ▸ hb₂ + field_simp [ha₂] linear_combination (W'.a₄ * W'.a₂ ^ 2 + W'.a₄ ^ 2) * CharP.cast_eq_zero F 3 theorem toCharThreeNF_spec_of_b₂_eq_zero (hb₂ : W.b₂ = 0) : diff --git a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean index 347525ce0df81..0ebab41ba1c2a 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean @@ -401,10 +401,10 @@ def extendMiddle (c : OrderedFinpartition n) (k : Fin c.length) : OrderedFinpart rcases eq_or_ne (c.index i) k with rfl | hi · have A : update c.partSize (c.index i) (c.partSize (c.index i) + 1) (c.index i) = c.partSize (c.index i) + 1 := by simp - exact ⟨c.index i, cast A.symm (succ (c.invEmbedding i)), by simp⟩ + exact ⟨c.index i, (succ (c.invEmbedding i)).cast A.symm , by simp⟩ · have A : update c.partSize k (c.partSize k + 1) (c.index i) = c.partSize (c.index i) := by simp [hi] - exact ⟨c.index i, cast A.symm (c.invEmbedding i), by simp [hi]⟩ + exact ⟨c.index i, (c.invEmbedding i).cast A.symm, by simp [hi]⟩ lemma index_extendMiddle_zero (c : OrderedFinpartition n) (i : Fin c.length) : (c.extendMiddle i).index 0 = i := by @@ -678,8 +678,8 @@ def extendEquiv (n : ℕ) : · simp only [↓reduceDIte, comp_apply] rcases eq_or_ne j 0 with rfl | hj · simpa using c.emb_zero - · let j' := Fin.pred (cast B.symm j) (by simpa using hj) - have : j = cast B (succ j') := by simp [j'] + · let j' := Fin.pred (j.cast B.symm) (by simpa using hj) + have : j = (succ j').cast B := by simp [j'] simp only [this, coe_cast, val_succ, cast_mk, cases_succ', comp_apply, succ_mk, Nat.succ_eq_add_one, succ_pred] rfl diff --git a/Mathlib/Analysis/Convex/Continuous.lean b/Mathlib/Analysis/Convex/Continuous.lean index fc4bc36368d8b..608b5a05cc364 100644 --- a/Mathlib/Analysis/Convex/Continuous.lean +++ b/Mathlib/Analysis/Convex/Continuous.lean @@ -51,7 +51,8 @@ lemma ConvexOn.lipschitzOnWith_of_abs_le (hf : ConvexOn ℝ (ball x₀ r) f) (h ε * (f x - f y) ≤ ‖x - y‖ * (f z - f x) := by rw [mul_sub, mul_sub, sub_le_sub_iff, ← add_mul] have h := hf.2 hy' hz (by positivity) (by positivity) hab - field_simp [← hxyz, a, b, ← mul_div_right_comm] at h + rw [← hxyz] at h + field_simp [a, b, ← mul_div_right_comm] at h rwa [← le_div_iff₀' (by positivity), add_comm (_ * _)] _ ≤ _ := by rw [sub_eq_add_neg (f _), two_mul] diff --git a/Mathlib/CategoryTheory/Galois/Decomposition.lean b/Mathlib/CategoryTheory/Galois/Decomposition.lean index 7b77438602714..0266fabf8bc16 100644 --- a/Mathlib/CategoryTheory/Galois/Decomposition.lean +++ b/Mathlib/CategoryTheory/Galois/Decomposition.lean @@ -154,12 +154,14 @@ lemma connected_component_unique {X A B : C} [IsConnected A] [IsConnected B] (a have : IsIso v := IsConnected.noTrivialComponent Y v hn use (asIso u).symm ≪≫ asIso v have hu : G.map u y = a := by - simp only [y, e, ← PreservesPullback.iso_hom_fst G, fiberPullbackEquiv, Iso.toEquiv_comp, - Equiv.symm_trans_apply, Iso.toEquiv_symm_fun, types_comp_apply, inv_hom_id_apply] + simp only [u, G, y, e, ← PreservesPullback.iso_hom_fst G, fiberPullbackEquiv, + Iso.toEquiv_comp, Equiv.symm_trans_apply, Iso.toEquiv_symm_fun, types_comp_apply, + inv_hom_id_apply] erw [Types.pullbackIsoPullback_inv_fst_apply (F.map i) (F.map j)] have hv : G.map v y = b := by - simp only [y, e, ← PreservesPullback.iso_hom_snd G, fiberPullbackEquiv, Iso.toEquiv_comp, - Equiv.symm_trans_apply, Iso.toEquiv_symm_fun, types_comp_apply, inv_hom_id_apply] + simp only [v, G, y, e, ← PreservesPullback.iso_hom_snd G, fiberPullbackEquiv, + Iso.toEquiv_comp, Equiv.symm_trans_apply, Iso.toEquiv_symm_fun, types_comp_apply, + inv_hom_id_apply] erw [Types.pullbackIsoPullback_inv_snd_apply (F.map i) (F.map j)] rw [← hu, ← hv] show (F.toPrefunctor.map u ≫ F.toPrefunctor.map _) y = F.toPrefunctor.map v y diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index 232291b05955f..e0c093182ad2c 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -1238,7 +1238,7 @@ theorem move_ok {p k₁ k₂ q s L₁ o L₂} {S : K' → List Γ'} (h₁ : k₁ swap · rw [Function.update_of_ne h₁.symm, List.reverseAux_nil] refine TransGen.head' rfl ?_ - simp only [TM2.step, Option.mem_def, TM2.stepAux, Option.elim, ne_eq] + rw [tr]; simp only [pop', TM2.stepAux] revert e; cases' S k₁ with a Sk <;> intro e · cases e rfl @@ -1248,7 +1248,7 @@ theorem move_ok {p k₁ k₂ q s L₁ o L₂} {S : K' → List Γ'} (h₁ : k₁ simp only [e] rfl · refine TransGen.head rfl ?_ - simp only [TM2.step, Option.mem_def, TM2.stepAux, Option.elim, ne_eq, List.reverseAux_cons] + rw [tr]; simp only [pop', Option.elim, TM2.stepAux, push'] cases' e₁ : S k₁ with a' Sk <;> rw [e₁, splitAtPred] at e · cases e cases e₂ : p a' <;> simp only [e₂, cond] at e @@ -1273,16 +1273,15 @@ theorem move₂_ok {p k₁ k₂ q s L₁ o L₂} {S : K' → List Γ'} (h₁ : k ⟨some q, none, update (update S k₁ (o.elim id List.cons L₂)) k₂ (L₁ ++ S k₂)⟩ := by refine (move_ok h₁.1 e).trans (TransGen.head rfl ?_) simp only [TM2.step, Option.mem_def, TM2.stepAux, id_eq, ne_eq, Option.elim] - cases o <;> simp only [Option.elim, id] - · simp only [TM2.stepAux, Option.isSome, cond_false] - convert move_ok h₁.2.1.symm (splitAtPred_false _) using 2 + cases o <;> simp only [Option.elim] <;> rw [tr] + <;> simp only [id, TM2.stepAux, Option.isSome, cond_true, cond_false] + · convert move_ok h₁.2.1.symm (splitAtPred_false _) using 2 simp only [Function.update_comm h₁.1, Function.update_idem] rw [show update S rev [] = S by rw [← h₂, Function.update_eq_self]] simp only [Function.update_of_ne h₁.2.2.symm, Function.update_of_ne h₁.2.1, Function.update_of_ne h₁.1.symm, List.reverseAux_eq, h₂, Function.update_self, List.append_nil, List.reverse_reverse] - · simp only [TM2.stepAux, Option.isSome, cond_true] - convert move_ok h₁.2.1.symm (splitAtPred_false _) using 2 + · convert move_ok h₁.2.1.symm (splitAtPred_false _) using 2 simp only [h₂, Function.update_comm h₁.1, List.reverseAux_eq, Function.update_self, List.append_nil, Function.update_idem] rw [show update S rev [] = S by rw [← h₂, Function.update_eq_self]] @@ -1293,7 +1292,7 @@ theorem clear_ok {p k q s L₁ o L₂} {S : K' → List Γ'} (e : splitAtPred p Reaches₁ (TM2.step tr) ⟨some (Λ'.clear p k q), s, S⟩ ⟨some q, o, update S k L₂⟩ := by induction' L₁ with a L₁ IH generalizing S s · refine TransGen.head' rfl ?_ - simp only [TM2.step, Option.mem_def, TM2.stepAux, Option.elim] + rw [tr]; simp only [pop', TM2.step, Option.mem_def, TM2.stepAux, Option.elim] revert e; cases' S k with a Sk <;> intro e · cases e rfl @@ -1303,7 +1302,7 @@ theorem clear_ok {p k q s L₁ o L₂} {S : K' → List Γ'} (e : splitAtPred p rcases e with ⟨e₁, e₂⟩ rw [e₁, e₂] · refine TransGen.head rfl ?_ - simp only [TM2.step, Option.mem_def, TM2.stepAux, Option.elim] + rw [tr]; simp only [pop', TM2.step, Option.mem_def, TM2.stepAux, Option.elim] cases' e₁ : S k with a' Sk <;> rw [e₁, splitAtPred] at e · cases e cases e₂ : p a' <;> simp only [e₂, cond] at e @@ -1322,9 +1321,10 @@ theorem copy_ok (q s a b c d) : · refine TransGen.single ?_ simp refine TransGen.head rfl ?_ + rw [tr] simp only [TM2.step, Option.mem_def, TM2.stepAux, elim_rev, List.head?_cons, Option.isSome_some, List.tail_cons, elim_update_rev, ne_eq, Function.update_of_ne, elim_main, elim_update_main, - elim_stack, elim_update_stack, cond_true, List.reverseAux_cons] + elim_stack, elim_update_stack, cond_true, List.reverseAux_cons, pop', push'] exact IH _ _ _ theorem trPosNum_natEnd : ∀ (n), ∀ x ∈ trPosNum n, natEnd x = false @@ -1358,6 +1358,7 @@ theorem head_main_ok {q s L} {c d : List Γ'} : (splitAtPred_eq _ _ (trNat L.headI) o (trList L.tail) (trNat_natEnd _) ?_)).trans (TransGen.head rfl (TransGen.head rfl ?_)) · cases L <;> simp [o] + rw [tr] simp only [TM2.step, Option.mem_def, TM2.stepAux, elim_update_main, elim_rev, elim_update_rev, Function.update_self, trList] rw [if_neg (show o ≠ some Γ'.consₗ by cases L <;> simp [o])] @@ -1375,6 +1376,7 @@ theorem head_stack_ok {q s L₁ L₂ L₃} : (move_ok (by decide) (splitAtPred_eq _ _ [] (some Γ'.consₗ) L₃ (by rintro _ ⟨⟩) ⟨rfl, rfl⟩)) (TransGen.head rfl (TransGen.head rfl ?_)) + rw [tr] simp only [TM2.step, Option.mem_def, TM2.stepAux, ite_true, id_eq, trList, List.nil_append, elim_update_stack, elim_rev, List.reverseAux_nil, elim_update_rev, Function.update_self, List.headI_nil, trNat_default] @@ -1428,7 +1430,8 @@ theorem succ_ok {q s n} {c d : List Γ'} : simp [PosNum.succ, trPosNum] rfl · refine ⟨l₁, _, some Γ'.bit0, rfl, TransGen.single ?_⟩ - simp only [TM2.step, TM2.stepAux, elim_main, elim_update_main, ne_eq, Function.update_of_ne, + simp only [TM2.step]; rw [tr] + simp only [TM2.stepAux, pop', elim_main, elim_update_main, ne_eq, Function.update_of_ne, elim_rev, elim_update_rev, Function.update_self, Option.mem_def, Option.some.injEq] rfl @@ -1445,11 +1448,9 @@ theorem pred_ok (q₁ q₂ s v) (c d : List Γ') : ∃ s', simp only [TM2.step, trList, trNat.eq_1, trNum, Nat.cast_succ, Num.add_one, Num.succ, List.tail_cons, List.headI_cons] cases' (n : Num) with a - · simp only [trPosNum, List.singleton_append, List.nil_append] + · simp only [trPosNum, Num.succ', List.singleton_append, List.nil_append] refine TransGen.head rfl ?_ - simp only [Option.mem_def, TM2.stepAux, elim_main, List.head?_cons, Option.some.injEq, - decide_false, List.tail_cons, elim_update_main, ne_eq, Function.update_of_ne, elim_rev, - elim_update_rev, natEnd, Function.update_self, cond_true, cond_false] + rw [tr]; simp only [pop', TM2.stepAux, cond_false] convert unrev_ok using 2 simp simp only [Num.succ'] @@ -1555,13 +1556,11 @@ theorem tr_ret_respects (k v s) : ∃ b₂, by_cases h : v.headI = 0 <;> simp only [h, ite_true, ite_false] at this ⊢ · obtain ⟨c, h₁, h₂⟩ := IH v.tail (trList v).head? refine ⟨c, h₁, TransGen.head rfl ?_⟩ - simp only [Option.mem_def, TM2.stepAux, trContStack, contStack, elim_main, this, cond_true, - elim_update_main] + rw [trCont, tr]; simp only [pop', TM2.stepAux, elim_main, this, elim_update_main] exact h₂ · obtain ⟨s', h₁, h₂⟩ := trNormal_respects f (Cont.fix f k) v.tail (some Γ'.cons) refine ⟨_, h₁, TransGen.head rfl <| TransGen.trans ?_ h₂⟩ - simp only [Option.mem_def, TM2.stepAux, elim_main, this.1, cond_false, elim_update_main, - trCont] + rw [trCont, tr]; simp only [pop', TM2.stepAux, elim_main, this.1] convert clear_ok (splitAtPred_eq _ _ (trNat v.headI).tail (some Γ'.cons) _ _ _) using 2 · simp convert rfl diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index 919e0be38ca9f..2d124e2707b94 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -1828,12 +1828,12 @@ theorem tr_respects : Respects (TM0.step M) (TM1.step (tr M)) fun a b ↦ trCfg cases' s with d a <;> rfl intro e refine TransGen.head ?_ (TransGen.head' this ?_) - · simp only [TM1.step, TM1.stepAux] + · simp only [TM1.step, TM1.stepAux, tr] rw [e] rfl cases e' : M q' _ · apply ReflTransGen.single - simp only [TM1.step, TM1.stepAux] + simp only [TM1.step, TM1.stepAux, tr] rw [e'] rfl · rfl @@ -1926,7 +1926,6 @@ section variable [DecidableEq K] /-- The step function for the TM2 model. -/ -@[simp] def stepAux : Stmt₂ → σ → (∀ k, List (Γ k)) → Cfg₂ | push k f q, v, S => stepAux q v (update S k (f v :: S k)) | peek k f q, v, S => stepAux q (f v (S k).head?) S @@ -1937,11 +1936,13 @@ def stepAux : Stmt₂ → σ → (∀ k, List (Γ k)) → Cfg₂ | halt, v, S => ⟨none, v, S⟩ /-- The step function for the TM2 model. -/ -@[simp] def step (M : Λ → Stmt₂) : Cfg₂ → Option Cfg₂ | ⟨none, _, _⟩ => none | ⟨some l, v, S⟩ => some (stepAux (M l) v S) +attribute [simp] stepAux.eq_1 stepAux.eq_2 stepAux.eq_3 + stepAux.eq_4 stepAux.eq_5 stepAux.eq_6 stepAux.eq_7 step.eq_1 step.eq_2 + /-- The (reflexive) reachability relation for the TM2 model. -/ def Reaches (M : Λ → Stmt₂) : Cfg₂ → Cfg₂ → Prop := ReflTransGen fun a b ↦ b ∈ step M a @@ -2292,7 +2293,7 @@ noncomputable def trStmts₁ : Stmt₂ → Finset Λ'₂₁ 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₁] + cases s <;> simp only [trStmts₁, stRun] theorem tr_respects_aux₂ [DecidableEq K] {k : K} {q : Stmt₂₁} {v : σ} {S : ∀ k, List (Γ k)} {L : ListBlank (∀ k, Option (Γ k))} diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 11123477cd16a..6056e283f70cc 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -269,7 +269,7 @@ This one instead uses a `NeZero n` typeclass hypothesis. theorem pos_iff_ne_zero' [NeZero n] (a : Fin n) : 0 < a ↔ a ≠ 0 := by rw [← val_fin_lt, val_zero', Nat.pos_iff_ne_zero, Ne, Ne, Fin.ext_iff, val_zero'] -@[simp] lemma cast_eq_self (a : Fin n) : cast rfl a = a := rfl +@[simp] lemma cast_eq_self (a : Fin n) : a.cast rfl = a := rfl @[simp] theorem cast_eq_zero {k l : ℕ} [NeZero k] [NeZero l] (h : k = l) (x : Fin k) : Fin.cast h x = 0 ↔ x = 0 := by simp [← val_eq_val] @@ -299,7 +299,7 @@ theorem revPerm_symm : (@revPerm n).symm = revPerm := rfl theorem cast_rev (i : Fin n) (h : n = m) : - cast h i.rev = (i.cast h).rev := by + i.rev.cast h = (i.cast h).rev := by subst h; simp theorem rev_eq_iff {i j : Fin n} : rev i = j ↔ i = rev j := by @@ -619,23 +619,23 @@ theorem coe_of_injective_castLE_symm {n k : ℕ} (h : n ≤ k) (i : Fin k) (hi) rw [← coe_castLE h] exact congr_arg Fin.val (Equiv.apply_ofInjective_symm _ _) -theorem leftInverse_cast (eq : n = m) : LeftInverse (cast eq.symm) (cast eq) := +theorem leftInverse_cast (eq : n = m) : LeftInverse (Fin.cast eq.symm) (Fin.cast eq) := fun _ => rfl -theorem rightInverse_cast (eq : n = m) : RightInverse (cast eq.symm) (cast eq) := +theorem rightInverse_cast (eq : n = m) : RightInverse (Fin.cast eq.symm) (Fin.cast eq) := fun _ => rfl -theorem cast_lt_cast (eq : n = m) {a b : Fin n} : cast eq a < cast eq b ↔ a < b := +theorem cast_lt_cast (eq : n = m) {a b : Fin n} : a.cast eq < b.cast eq ↔ a < b := Iff.rfl -theorem cast_le_cast (eq : n = m) {a b : Fin n} : cast eq a ≤ cast eq b ↔ a ≤ b := +theorem cast_le_cast (eq : n = m) {a b : Fin n} : a.cast eq ≤ b.cast eq ↔ a ≤ b := Iff.rfl /-- The 'identity' equivalence between `Fin m` and `Fin n` when `m = n`. -/ @[simps] def _root_.finCongr (eq : n = m) : Fin n ≃ Fin m where - toFun := cast eq - invFun := cast eq.symm + toFun := Fin.cast eq + invFun := Fin.cast eq.symm left_inv := leftInverse_cast eq right_inv := rightInverse_cast eq @@ -656,12 +656,12 @@ a generic theorem about `cast`. -/ lemma _root_.finCongr_eq_equivCast (h : n = m) : finCongr h = .cast (h ▸ rfl) := by subst h; simp @[simp] -theorem cast_zero {n' : ℕ} [NeZero n] {h : n = n'} : cast h (0 : Fin n) = +theorem cast_zero {n' : ℕ} [NeZero n] {h : n = n'} : (0 : Fin n).cast h = by { haveI : NeZero n' := by {rw [← h]; infer_instance}; exact 0} := rfl /-- While in many cases `Fin.cast` is better than `Equiv.cast`/`cast`, sometimes we want to apply a generic theorem about `cast`. -/ -theorem cast_eq_cast (h : n = m) : (cast h : Fin n → Fin m) = _root_.cast (h ▸ rfl) := by +theorem cast_eq_cast (h : n = m) : (Fin.cast h : Fin n → Fin m) = _root_.cast (h ▸ rfl) := by subst h ext rfl diff --git a/Mathlib/Data/Fin/Tuple/Basic.lean b/Mathlib/Data/Fin/Tuple/Basic.lean index 72ddaaa0b028e..3a31ba890f95d 100644 --- a/Mathlib/Data/Fin/Tuple/Basic.lean +++ b/Mathlib/Data/Fin/Tuple/Basic.lean @@ -395,7 +395,7 @@ theorem cons_eq_append (x : α) (xs : Fin n → α) : subst h; simp lemma append_rev {m n} (xs : Fin m → α) (ys : Fin n → α) (i : Fin (m + n)) : - append xs ys (rev i) = append (ys ∘ rev) (xs ∘ rev) (cast (Nat.add_comm ..) i) := by + append xs ys (rev i) = append (ys ∘ rev) (xs ∘ rev) (i.cast (Nat.add_comm ..)) := by rcases rev_surjective i with ⟨i, rfl⟩ rw [rev_rev] induction i using Fin.addCases @@ -403,7 +403,7 @@ lemma append_rev {m n} (xs : Fin m → α) (ys : Fin n → α) (i : Fin (m + n)) · simp [cast_rev, rev_addNat] lemma append_comp_rev {m n} (xs : Fin m → α) (ys : Fin n → α) : - append xs ys ∘ rev = append (ys ∘ rev) (xs ∘ rev) ∘ cast (Nat.add_comm ..) := + append xs ys ∘ rev = append (ys ∘ rev) (xs ∘ rev) ∘ Fin.cast (Nat.add_comm ..) := funext <| append_rev xs ys theorem append_castAdd_natAdd {f : Fin (m + n) → α} : @@ -430,11 +430,11 @@ theorem repeat_apply (a : Fin n → α) (i : Fin (m * n)) : @[simp] theorem repeat_zero (a : Fin n → α) : - Fin.repeat 0 a = Fin.elim0 ∘ cast (Nat.zero_mul _) := - funext fun x => (cast (Nat.zero_mul _) x).elim0 + Fin.repeat 0 a = Fin.elim0 ∘ Fin.cast (Nat.zero_mul _) := + funext fun x => (x.cast (Nat.zero_mul _)).elim0 @[simp] -theorem repeat_one (a : Fin n → α) : Fin.repeat 1 a = a ∘ cast (Nat.one_mul _) := by +theorem repeat_one (a : Fin n → α) : Fin.repeat 1 a = a ∘ Fin.cast (Nat.one_mul _) := by generalize_proofs h apply funext rw [(Fin.rightInverse_cast h.symm).surjective.forall] @@ -443,7 +443,7 @@ theorem repeat_one (a : Fin n → α) : Fin.repeat 1 a = a ∘ cast (Nat.one_mul theorem repeat_succ (a : Fin n → α) (m : ℕ) : Fin.repeat m.succ a = - append a (Fin.repeat m a) ∘ cast ((Nat.succ_mul _ _).trans (Nat.add_comm ..)) := by + append a (Fin.repeat m a) ∘ Fin.cast ((Nat.succ_mul _ _).trans (Nat.add_comm ..)) := by generalize_proofs h apply funext rw [(Fin.rightInverse_cast h.symm).surjective.forall] @@ -453,7 +453,7 @@ theorem repeat_succ (a : Fin n → α) (m : ℕ) : @[simp] theorem repeat_add (a : Fin n → α) (m₁ m₂ : ℕ) : Fin.repeat (m₁ + m₂) a = - append (Fin.repeat m₁ a) (Fin.repeat m₂ a) ∘ cast (Nat.add_mul ..) := by + append (Fin.repeat m₁ a) (Fin.repeat m₂ a) ∘ Fin.cast (Nat.add_mul ..) := by generalize_proofs h apply funext rw [(Fin.rightInverse_cast h.symm).surjective.forall] diff --git a/Mathlib/Data/Fin/Tuple/Take.lean b/Mathlib/Data/Fin/Tuple/Take.lean index 232622468b153..56346dbf9bff5 100644 --- a/Mathlib/Data/Fin/Tuple/Take.lean +++ b/Mathlib/Data/Fin/Tuple/Take.lean @@ -132,7 +132,7 @@ theorem take_addCases_right {n' : ℕ} {motive : Fin (n + n') → Sort*} (m : by_cases h' : i < n · simp only [h', ↓reduceDIte] congr - · simp only [h', ↓reduceDIte, subNat, castLE, cast, eqRec_eq_cast] + · simp only [h', ↓reduceDIte, subNat, castLE, Fin.cast, eqRec_eq_cast] /-- Version of `take_addCases_right` that specializes `addCases` to `append`. -/ theorem take_append_right {n' : ℕ} {α : Sort*} (m : ℕ) (h : m ≤ n') (u : (i : Fin n) → α) diff --git a/Mathlib/Data/Int/Bitwise.lean b/Mathlib/Data/Int/Bitwise.lean index e7ab67fafc45c..3e4f9127bad67 100644 --- a/Mathlib/Data/Int/Bitwise.lean +++ b/Mathlib/Data/Int/Bitwise.lean @@ -399,8 +399,6 @@ theorem shiftRight_add' : ∀ (m : ℤ) (n k : ℕ), m >>> (n + k : ℤ) = (m >> /-! ### bitwise ops -/ -attribute [local simp] Int.zero_div - theorem shiftLeft_add : ∀ (m : ℤ) (n : ℕ) (k : ℤ), m <<< (n + k) = (m <<< (n : ℤ)) <<< k | (m : ℕ), n, (k : ℕ) => congr_arg ofNat (by simp [Nat.shiftLeft_eq, Nat.pow_add, mul_assoc]) diff --git a/Mathlib/Data/Int/Lemmas.lean b/Mathlib/Data/Int/Lemmas.lean index f98d34103dd0d..c51bbeb9c0106 100644 --- a/Mathlib/Data/Int/Lemmas.lean +++ b/Mathlib/Data/Int/Lemmas.lean @@ -112,9 +112,6 @@ theorem toNat_of_nonpos : ∀ {z : ℤ}, z ≤ 0 → z.toNat = 0 This lemma is orphaned from `Data.Int.Bitwise` as it also requires material from `Data.Int.Order`. -/ - -attribute [local simp] Int.zero_div - @[simp] theorem div2_bit (b n) : div2 (bit b n) = n := by rw [bit_val, div2_val, add_comm, Int.add_mul_ediv_left, (_ : (_ / 2 : ℤ) = 0), zero_add] diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index dfd2fc1bb362c..4c99fc84425e5 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -211,9 +211,6 @@ theorem eq_replicate_length {a : α} : ∀ {l : List α}, l = replicate l.length theorem replicate_add (m n) (a : α) : replicate (m + n) a = replicate m a ++ replicate n a := by rw [append_replicate_replicate] -theorem replicate_succ' (n) (a : α) : replicate (n + 1) a = replicate n a ++ [a] := - replicate_add n 1 a - theorem replicate_subset_singleton (n) (a : α) : replicate n a ⊆ [a] := fun _ h => mem_singleton.2 (eq_of_mem_replicate h) @@ -300,7 +297,7 @@ theorem getLast_append' (l₁ l₂ : List α) (h : l₂ ≠ []) : | nil => simp | cons _ _ ih => 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 := by +theorem getLast_concat' {a : α} (l : List α) : getLast (concat l a) (by simp) = a := by simp @[simp] @@ -487,10 +484,6 @@ theorem forall_mem_iff_getElem {l : List α} {p : α → Prop} : (∀ x ∈ l, p x) ↔ ∀ (i : ℕ) (_ : i < l.length), p l[i] := by simp [mem_iff_getElem, @forall_swap α] -theorem getElem_cons {l : List α} {a : α} {n : ℕ} (h : n < (a :: l).length) : - (a :: l)[n] = if hn : n = 0 then a else l[n - 1]'(by rw [length_cons] at h; omega) := by - cases n <;> simp - theorem get_tail (l : List α) (i) (h : i < l.tail.length) (h' : i + 1 < l.length := (by simp only [length_tail] at h; omega)) : l.tail.get ⟨i, h⟩ = l.get ⟨i + 1, h'⟩ := by diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index 74ed472f6b59f..f209731a712ce 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -421,7 +421,7 @@ theorem Chain'.cons_of_le [LinearOrder α] {a : α} {as m : List α} cases as with | nil => simp only [le_iff_lt_or_eq, reduceCtorEq, or_false] at hmas - exact (List.Lex.not_nil_right (·<·) _ hmas).elim + exact (List.not_lt_nil _ hmas).elim | cons a' as => rw [List.chain'_cons] at ha refine gt_of_gt_of_ge ha.1 ?_ @@ -431,7 +431,7 @@ theorem Chain'.cons_of_le [LinearOrder α] {a : α} {as m : List α} rw [← not_le] at hmas apply hmas apply le_of_lt - exact (List.lt_iff_lex_lt _ _).mp (List.lt.head _ _ hh) + exact (List.lt_iff_lex_lt _ _).mp (List.Lex.rel hh) · simp_all only [List.cons.injEq, le_refl] lemma Chain'.chain {α : Type*} {R : α → α → Prop} {l : List α} {v : α} diff --git a/Mathlib/Data/List/Lemmas.lean b/Mathlib/Data/List/Lemmas.lean index ef62ad8d898a9..f9ef0d4e07c18 100644 --- a/Mathlib/Data/List/Lemmas.lean +++ b/Mathlib/Data/List/Lemmas.lean @@ -17,17 +17,9 @@ variable {α β γ : Type*} namespace List -@[simp] -theorem length_flatMap (l : List α) (f : α → List β) : - length (List.flatMap l f) = sum (map (length ∘ f) l) := by - rw [List.flatMap, length_flatten, map_map] - -lemma countP_flatMap (p : β → Bool) (l : List α) (f : α → List β) : - countP p (l.flatMap f) = sum (map (countP p ∘ f) l) := by - rw [List.flatMap, countP_flatten, map_map] - -lemma count_flatMap [BEq β] (l : List α) (f : α → List β) (x : β) : - count x (l.flatMap f) = sum (map (count x ∘ f) l) := countP_flatMap _ _ _ +set_option linter.deprecated false in +@[simp, deprecated "No deprecation message was provided." (since := "2024-10-17")] +lemma Nat.sum_eq_listSum (l : List ℕ) : Nat.sum l = l.sum := rfl @[deprecated (since := "2024-08-20")] alias getElem_reverse' := getElem_reverse diff --git a/Mathlib/Data/List/Lex.lean b/Mathlib/Data/List/Lex.lean index 69566d9aa5d91..4c166bc4b024a 100644 --- a/Mathlib/Data/List/Lex.lean +++ b/Mathlib/Data/List/Lex.lean @@ -35,24 +35,13 @@ variable {α : Type u} /-! ### lexicographic ordering -/ - -/-- Given a strict order `<` on `α`, the lexicographic strict order on `List α`, for which -`[a0, ..., an] < [b0, ..., b_k]` if `a0 < b0` or `a0 = b0` and `[a1, ..., an] < [b1, ..., bk]`. -The definition is given for any relation `r`, not only strict orders. -/ -inductive Lex (r : α → α → Prop) : List α → List α → Prop - | nil {a l} : Lex r [] (a :: l) - | cons {a l₁ l₂} (h : Lex r l₁ l₂) : Lex r (a :: l₁) (a :: l₂) - | rel {a₁ l₁ a₂ l₂} (h : r a₁ a₂) : Lex r (a₁ :: l₁) (a₂ :: l₂) - namespace Lex theorem cons_iff {r : α → α → Prop} [IsIrrefl α r] {a l₁ l₂} : Lex r (a :: l₁) (a :: l₂) ↔ Lex r l₁ l₂ := ⟨fun h => by cases' h with _ _ _ _ _ h _ _ _ _ h; exacts [h, (irrefl_of r a h).elim], Lex.cons⟩ -@[simp] -theorem not_nil_right (r : α → α → Prop) (l : List α) : ¬Lex r l [] := - nofun +@[deprecated (since := "2024-12-21")] alias not_nil_right := not_lex_nil theorem nil_left_or_eq_nil {r : α → α → Prop} (l : List α) : List.Lex r [] l ∨ l = [] := match l with @@ -158,9 +147,6 @@ end Lex instance LT' [LT α] : LT (List α) := ⟨Lex (· < ·)⟩ -theorem nil_lt_cons [LT α] (a : α) (l : List α) : [] < a :: l := - Lex.nil - instance [LinearOrder α] : LinearOrder (List α) := linearOrderOfSTO (Lex (· < ·)) @@ -168,25 +154,8 @@ instance [LinearOrder α] : LinearOrder (List α) := instance LE' [LinearOrder α] : LE (List α) := Preorder.toLE -theorem lt_iff_lex_lt [LinearOrder α] (l l' : List α) : lt l l' ↔ Lex (· < ·) l l' := by - constructor <;> - intro h - · induction h with - | nil b bs => exact Lex.nil - | @head a as b bs hab => apply Lex.rel; assumption - | @tail a as b bs hab hba _ ih => - have heq : a = b := _root_.le_antisymm (le_of_not_lt hba) (le_of_not_lt hab) - subst b; apply Lex.cons; assumption - · induction h with - | @nil a as => apply lt.nil - | @cons a as bs _ ih => apply lt.tail <;> simp [ih] - | @rel a as b bs h => apply lt.head; assumption - -@[simp] -theorem nil_le {α} [LinearOrder α] {l : List α} : [] ≤ l := - match l with - | [] => le_rfl - | _ :: _ => le_of_lt <| nil_lt_cons _ _ +theorem lt_iff_lex_lt [LinearOrder α] (l l' : List α) : List.lt l l' ↔ Lex (· < ·) l l' := by + rw [List.lt] theorem head_le_of_lt [Preorder α] {a a' : α} {l l' : List α} (h : (a' :: l') < (a :: l)) : a' ≤ a := diff --git a/Mathlib/Data/List/MinMax.lean b/Mathlib/Data/List/MinMax.lean index cb9db4834736d..71564ef6cb848 100644 --- a/Mathlib/Data/List/MinMax.lean +++ b/Mathlib/Data/List/MinMax.lean @@ -449,7 +449,7 @@ lemma getD_max?_eq_unbot'_maximum (l : List α) (d : α) : cases hz : l.max? with | none => simp [List.max?_eq_none_iff.mp hz] at hy | some z => - have : Std.Antisymm (α := α) (· ≤ ·) := ⟨_root_.le_antisymm⟩ + have : Std.Antisymm (α := α) (· ≤ ·) := ⟨fun _ _ => _root_.le_antisymm⟩ rw [List.max?_eq_some_iff] at hz · rw [Option.getD_some] exact _root_.le_antisymm (hy.right _ hz.left) (hz.right _ hy.left) diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index 49f6dd007cd72..bdd3205fa2f21 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -89,7 +89,7 @@ theorem map_permutationsAux2' {α' β'} (g : α → α') (g' : β → β') (t : induction' ys with ys_hd _ ys_ih generalizing f f' · simp · simp only [map, permutationsAux2_snd_cons, cons_append, cons.injEq] - rw [ys_ih, permutationsAux2_fst] + rw [ys_ih] · refine ⟨?_, rfl⟩ simp only [← map_cons, ← map_append]; apply H · intro a; apply H @@ -268,7 +268,7 @@ theorem length_permutationsAux : refine permutationsAux.rec (by simp) ?_ intro t ts is IH1 IH2 have IH2 : length (permutationsAux is nil) + 1 = is.length ! := by simpa using IH2 - simp only [factorial, Nat.mul_comm, add_eq] at IH1 + simp only [length, factorial, Nat.mul_comm, add_eq] at IH1 rw [permutationsAux_cons, length_foldr_permutationsAux2' _ _ _ _ _ fun l m => (perm_of_mem_permutations m).length_eq, permutations, length, length, IH2, Nat.succ_add, Nat.factorial_succ, Nat.mul_comm (_ + 1), @@ -387,7 +387,7 @@ theorem getElem_permutations'Aux (s : List α) (x : α) (n : ℕ) (hn : n < length (permutations'Aux x s)) : (permutations'Aux x s)[n] = s.insertIdx n x := by induction' s with y s IH generalizing n - · simp only [length, Nat.zero_add, Nat.lt_one_iff] at hn + · simp only [permutations'Aux, length, Nat.zero_add, lt_one_iff] at hn simp [hn] · cases n · simp [get] @@ -509,7 +509,7 @@ theorem nodup_permutations (s : List α) (hs : Nodup s) : Nodup s.permutations : 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] + induction s <;> simp [permutations, permutationsAux.rec] @[simp] theorem nodup_permutations_iff {s : List α} : Nodup s.permutations ↔ Nodup s := by diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index cf9e742f04bfc..c321ba243c9b5 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -80,12 +80,9 @@ theorem binaryRec_of_ne_zero {C : Nat → Sort*} (z : C 0) (f : ∀ b n, C n → lemma bitwise_bit {f : Bool → Bool → Bool} (h : f false false = false := by rfl) (a m b n) : bitwise f (bit a m) (bit b n) = bit (f a b) (bitwise f m n) := by conv_lhs => unfold bitwise - #adaptation_note /-- nightly-2024-03-16: simp was - -- simp (config := { unfoldPartialApp := true }) only [bit, bit1, bit0, Bool.cond_eq_ite] -/ simp only [bit, ite_apply, Bool.cond_eq_ite] - have h2 x : (x + x + 1) % 2 = 1 := by rw [← two_mul, add_comm]; apply add_mul_mod_self_left have h4 x : (x + x + 1) / 2 = x := by rw [← two_mul, add_comm]; simp [add_mul_div_left] - cases a <;> cases b <;> simp [h2, h4] <;> split_ifs + cases a <;> cases b <;> simp [h4] <;> split_ifs <;> simp_all +decide [two_mul] lemma bit_mod_two_eq_zero_iff (a x) : diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 1e19ccb7c114d..5f52c3eb53cf2 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -978,15 +978,12 @@ lemma mod_two_ne_zero : n % 2 ≠ 0 ↔ n % 2 = 1 := mod_two_not_eq_zero lemma div_mod_eq_mod_mul_div (a b c : ℕ) : a / b % c = a % (b * c) / b := (mod_mul_right_div_self a b c).symm -/-- Variant of `Nat.lt_div_iff_mul_lt` (added in Lean 4.16) that assumes `d ∣ n`. -/ +/-- Variant of `Nat.lt_div_iff_mul_lt` that assumes `d ∣ n`. -/ protected lemma lt_div_iff_mul_lt' (hdn : d ∣ n) (a : ℕ) : a < n / d ↔ d * a < n := by obtain rfl | hd := d.eq_zero_or_pos · simp [Nat.zero_dvd.1 hdn] · rw [← Nat.mul_lt_mul_left hd, ← Nat.eq_mul_of_div_eq_right hdn rfl] -@[deprecated Nat.lt_div_iff_mul_lt' (since := "2024-12-29")] -protected alias lt_div_iff_mul_lt := Nat.lt_div_iff_mul_lt' - lemma mul_div_eq_iff_dvd {n d : ℕ} : d * (n / d) = n ↔ d ∣ n := calc d * (n / d) = n ↔ d * (n / d) = d * (n / d) + (n % d) := by rw [div_add_mod] diff --git a/Mathlib/Data/Num/Basic.lean b/Mathlib/Data/Num/Basic.lean index 54c9466d8c312..25e1a56167888 100644 --- a/Mathlib/Data/Num/Basic.lean +++ b/Mathlib/Data/Num/Basic.lean @@ -138,7 +138,7 @@ def ofNatSucc : ℕ → PosNum def ofNat (n : ℕ) : PosNum := ofNatSucc (Nat.pred n) -instance {n : ℕ} : OfNat PosNum (n + 1) where +instance (priority := low) {n : ℕ} : OfNat PosNum (n + 1) where ofNat := ofNat (n + 1) open Ordering diff --git a/Mathlib/Data/QPF/Multivariate/Constructions/Fix.lean b/Mathlib/Data/QPF/Multivariate/Constructions/Fix.lean index 284cbecc9b925..02db3344da144 100644 --- a/Mathlib/Data/QPF/Multivariate/Constructions/Fix.lean +++ b/Mathlib/Data/QPF/Multivariate/Constructions/Fix.lean @@ -309,7 +309,7 @@ def Fix.drec {β : Fix F α → Type u} rw [Fix.rec_eq] dsimp simp? [appendFun_id_id] at ih says - simp only [appendFun_id_id, MvFunctor.id_map] at ih + simp only [appendFun_id_id, MvFunctor.id_map, y] at ih congr conv => rhs diff --git a/Mathlib/Data/String/Basic.lean b/Mathlib/Data/String/Basic.lean index b4a8b44ddfbf8..7bd4035fff1aa 100644 --- a/Mathlib/Data/String/Basic.lean +++ b/Mathlib/Data/String/Basic.lean @@ -28,10 +28,12 @@ def ltb (s₁ s₂ : Iterator) : Bool := else true else false +/-- This overrides an instance in core Lean. -/ instance LT' : LT String := ⟨fun s₁ s₂ ↦ ltb s₁.iter s₂.iter⟩ -instance decidableLT : DecidableRel (α := String) (· < ·) := by +/-- This instance has a prime to avoid the name of the corresponding instance in core Lean. -/ +instance decidableLT' : DecidableRel (α := String) (· < ·) := by simp only [LT'] infer_instance -- short-circuit type class inference diff --git a/Mathlib/Data/Vector/Basic.lean b/Mathlib/Data/Vector/Basic.lean index 8e81475a4c492..7f7f4602004bc 100644 --- a/Mathlib/Data/Vector/Basic.lean +++ b/Mathlib/Data/Vector/Basic.lean @@ -317,7 +317,7 @@ This lemma is the `cons` version of `scanl_get`. @[simp] theorem scanl_cons (x : α) : scanl f b (x ::ᵥ v) = b ::ᵥ scanl f (f b x) v := by simp only [scanl, toList_cons, List.scanl]; dsimp - simp only [cons]; rfl + simp only [cons] /-- The underlying `List` of a `Vector` after a `scanl` is the `List.scanl` of the underlying `List` of the original `Vector`. @@ -351,8 +351,7 @@ theorem scanl_head : (scanl f b v).head = b := by · 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, toList_scanl, toList_cons, List.scanl, Fin.val_zero, - List.get] + simp [← get_zero, get_eq_get_toList] /-- For an index `i : Fin n`, the nth element of `scanl` of a vector `v : Vector α n` at `i.succ`, is equal to the application diff --git a/Mathlib/Deprecated/LazyList.lean b/Mathlib/Deprecated/LazyList.lean index 04f7bd5f25dc5..80dbb7f596a7a 100644 --- a/Mathlib/Deprecated/LazyList.lean +++ b/Mathlib/Deprecated/LazyList.lean @@ -102,7 +102,7 @@ instance : LawfulMonad LazyList := LawfulMonad.mk' (bind_pure_comp := by intro _ _ f xs simp only [bind, Functor.map, pure, singleton] - induction xs using LazyList.traverse.induct (m := @Id) (f := f) with + induction xs using LazyList.traverse.induct with | case1 => simp only [Thunk.pure, LazyList.bind, LazyList.traverse, Id.pure_eq] | case2 _ _ ih => diff --git a/Mathlib/FieldTheory/Finite/GaloisField.lean b/Mathlib/FieldTheory/Finite/GaloisField.lean index f2ef367ab073b..e4a08ea36a937 100644 --- a/Mathlib/FieldTheory/Finite/GaloisField.lean +++ b/Mathlib/FieldTheory/Finite/GaloisField.lean @@ -112,7 +112,9 @@ theorem finrank {n} (h : n ≠ 0) : Module.finrank (ZMod p) (GaloisField p n) = intro x hx -- We discharge the `p = 0` separately, to avoid typeclass issues on `ZMod p`. cases p; cases hp - refine Subring.closure_induction ?_ ?_ ?_ ?_ ?_ ?_ hx <;> simp_rw [mem_rootSet_of_ne aux] + simp only [g_poly] at aux + refine Subring.closure_induction ?_ ?_ ?_ ?_ ?_ ?_ hx + <;> simp_rw [mem_rootSet_of_ne aux] · rintro x (⟨r, rfl⟩ | hx) · simp only [g_poly, map_sub, map_pow, aeval_X] rw [← map_pow, ZMod.pow_card_pow, sub_self] diff --git a/Mathlib/Lean/Expr/Basic.lean b/Mathlib/Lean/Expr/Basic.lean index 4606e2f6b5692..e00f5ad914dd0 100644 --- a/Mathlib/Lean/Expr/Basic.lean +++ b/Mathlib/Lean/Expr/Basic.lean @@ -199,7 +199,7 @@ def eraseProofs (e : Expr) : MetaM Expr := Meta.transform (skipConstInApp := true) e (pre := fun e => do if (← Meta.isProof e) then - return .continue (← mkSyntheticSorry (← inferType e)) + return .continue (← mkSorry (← inferType e) true) else return .continue) diff --git a/Mathlib/Lean/Meta/RefinedDiscrTree/Encode.lean b/Mathlib/Lean/Meta/RefinedDiscrTree/Encode.lean index 18f7876c84b89..2abc81e56ab46 100644 --- a/Mathlib/Lean/Meta/RefinedDiscrTree/Encode.lean +++ b/Mathlib/Lean/Meta/RefinedDiscrTree/Encode.lean @@ -100,7 +100,7 @@ where 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 + unless ← go as[i] (bs[i]'(h ▸ g.2.1)) do return false return true else diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index c9f988704f257..b8ca669106157 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -794,7 +794,7 @@ theorem _root_.mem_span_of_iInf_ker_le_ker [Finite ι] {L : ι → E →ₗ[𝕜 let L' i : (E ⧸ p) →ₗ[𝕜] 𝕜 := p.liftQ (L i) (iInf_le _ i) let K' : (E ⧸ p) →ₗ[𝕜] 𝕜 := p.liftQ K h have : ⨅ i, ker (L' i) ≤ ker K' := by - simp_rw [← ker_pi, L', pi_liftQ_eq_liftQ_pi, ker_liftQ_eq_bot' p φ p_eq] + simp_rw +zetaDelta [← ker_pi, pi_liftQ_eq_liftQ_pi, ker_liftQ_eq_bot' p φ p_eq] exact bot_le obtain ⟨c, hK'⟩ := (mem_span_range_iff_exists_fun 𝕜).1 (FiniteDimensional.mem_span_of_iInf_ker_le_ker this) diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index aa5849bf7df41..973a47eec399b 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -1366,10 +1366,10 @@ lemma exists_linearIndependent' (v : ι → V) : have hs {i : ι} (hi : i ∈ s) : v i ∈ t := by obtain ⟨a, rfl⟩ := hi; simp [hf] let f' (a : s) : t := ⟨v a.val, hs a.property⟩ refine ⟨s, Subtype.val, Subtype.val_injective, hsp.symm ▸ by congr; aesop, ?_⟩ - · rw [← show Subtype.val ∘ f' = v ∘ Subtype.val by ext; simp] + · rw [← show Subtype.val ∘ f' = v ∘ Subtype.val by ext; simp [f']] apply hli.comp rintro ⟨i, x, rfl⟩ ⟨j, y, rfl⟩ hij - simp only [Subtype.ext_iff, hf] at hij + simp only [Subtype.ext_iff, hf, f'] at hij simp [hij] variable {K t} diff --git a/Mathlib/LinearAlgebra/Matrix/SemiringInverse.lean b/Mathlib/LinearAlgebra/Matrix/SemiringInverse.lean index e6d05239e6324..72c3fbee941ce 100644 --- a/Mathlib/LinearAlgebra/Matrix/SemiringInverse.lean +++ b/Mathlib/LinearAlgebra/Matrix/SemiringInverse.lean @@ -122,7 +122,7 @@ theorem mul_adjp_apply_ne (h : i ≠ j) : (A * adjp 1 A) i j = (A * adjp (-1) A) have key : ({j}ᶜ : Finset n) = disjUnion ({i} : Finset n) ({i, j} : Finset n)ᶜ (by simp) := by rw [singleton_disjUnion, cons_eq_insert, compl_insert, insert_erase] rwa [mem_compl, mem_singleton] - simp_rw [key, prod_disjUnion, prod_singleton, Perm.mul_apply, swap_apply_left, ← mul_assoc] + simp_rw [key, prod_disjUnion, prod_singleton, f, Perm.mul_apply, swap_apply_left, ← mul_assoc] rw [mul_comm (A i x) (A i (σ i)), hp.2.2] refine congr_arg _ (prod_congr rfl fun x hx ↦ ?_) rw [mem_compl, mem_insert, mem_singleton, not_or] at hx diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index d63de1913084e..409b30a71b81e 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -478,7 +478,7 @@ than `forall_swap`. -/ theorem imp_forall_iff {α : Type*} {p : Prop} {q : α → Prop} : (p → ∀ x, q x) ↔ ∀ x, p → q x := forall_swap -@[simp] lemma imp_forall_iff_forall (A : Prop) (B : A → Prop) : +lemma imp_forall_iff_forall (A : Prop) (B : A → Prop) : (A → ∀ h : A, B h) ↔ ∀ h : A, B h := by by_cases h : A <;> simp [h] theorem exists_swap {p : α → β → Prop} : (∃ x y, p x y) ↔ ∃ y x, p x y := diff --git a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean index 7dbef7d066822..a88490002c8be 100644 --- a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean +++ b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean @@ -1193,7 +1193,7 @@ lemma integral_unitInterval_deriv_eq_sub [RCLike 𝕜] [NormedSpace 𝕜 E] [IsS simp only [one_smul] exact this.const_add z₀ convert (integral_eq_sub_of_hasDerivAt hderiv' hint) using 1 - · simp_rw [← integral_smul, Function.comp_apply] + · simp_rw [← integral_smul, Function.comp_apply, γ] · simp only [γ, Function.comp_apply, one_smul, zero_smul, add_zero] /-! diff --git a/Mathlib/MeasureTheory/Integral/Marginal.lean b/Mathlib/MeasureTheory/Integral/Marginal.lean index 483452c38db13..69a96847de886 100644 --- a/Mathlib/MeasureTheory/Integral/Marginal.lean +++ b/Mathlib/MeasureTheory/Integral/Marginal.lean @@ -124,7 +124,7 @@ theorem lmarginal_singleton (f : (∀ i, π i) → ℝ≥0∞) (i : δ) : = ∫⁻ (y : π (default : α)), f (updateFinset x {i} (e y)) ∂μ (default : α) := by simp_rw [lmarginal, measurePreserving_piUnique (fun j : ({i} : Finset δ) ↦ μ j) |>.symm _ - |>.lintegral_map_equiv] + |>.lintegral_map_equiv, e, α] _ = ∫⁻ xᵢ, f (Function.update x i xᵢ) ∂μ i := by simp [update_eq_updateFinset]; rfl variable {μ} in diff --git a/Mathlib/ModelTheory/Semantics.lean b/Mathlib/ModelTheory/Semantics.lean index 01e220cf30c65..26c5afe142981 100644 --- a/Mathlib/ModelTheory/Semantics.lean +++ b/Mathlib/ModelTheory/Semantics.lean @@ -323,7 +323,7 @@ theorem realize_iff : (φ.iff ψ).Realize v xs ↔ (φ.Realize v xs ↔ ψ.Reali simp only [BoundedFormula.iff, realize_inf, realize_imp, and_imp, ← iff_def] theorem realize_castLE_of_eq {m n : ℕ} (h : m = n) {h' : m ≤ n} {φ : L.BoundedFormula α m} - {v : α → M} {xs : Fin n → M} : (φ.castLE h').Realize v xs ↔ φ.Realize v (xs ∘ cast h) := by + {v : α → M} {xs : Fin n → M} : (φ.castLE h').Realize v xs ↔ φ.Realize v (xs ∘ Fin.cast h) := by subst h simp only [castLE_rfl, cast_refl, OrderIso.coe_refl, Function.comp_id] @@ -422,7 +422,7 @@ theorem realize_restrictFreeVar [DecidableEq α] {n : ℕ} {φ : L.BoundedFormul induction φ with | falsum => rfl | equal => - simp only [Realize, freeVarFinset.eq_2] + simp only [Realize, restrictFreeVar, freeVarFinset.eq_2] rw [realize_restrictVarLeft v' (by simp [hv']), realize_restrictVarLeft v' (by simp [hv'])] simp [Function.comp_apply] | rel => @@ -431,7 +431,7 @@ theorem realize_restrictFreeVar [DecidableEq α] {n : ℕ} {φ : L.BoundedFormul rw [realize_restrictVarLeft v' (by simp [hv'])] simp [Function.comp_apply] | imp _ _ ih1 ih2 => - simp only [Realize, freeVarFinset.eq_4] + simp only [Realize, restrictFreeVar, freeVarFinset.eq_4] rw [ih1, ih2] <;> simp [hv'] | all _ ih3 => simp only [restrictFreeVar, Realize] diff --git a/Mathlib/NumberTheory/Bernoulli.lean b/Mathlib/NumberTheory/Bernoulli.lean index 8e39a25fbcd57..b44c36b5a8f8e 100644 --- a/Mathlib/NumberTheory/Bernoulli.lean +++ b/Mathlib/NumberTheory/Bernoulli.lean @@ -291,7 +291,7 @@ theorem sum_range_pow (n p : ℕ) : simp only [f, exp_pow_eq_rescale_exp, rescale, one_div, coeff_mk, RingHom.coe_mk, coeff_exp, 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 + simp? at h says simp only [succ_eq_add_one, mem_range, f] at h rw [choose_eq_factorial_div_factorial h.le, eq_comm, div_eq_iff (hne q.succ), succ_eq_add_one, mul_assoc _ _ (q.succ ! : ℚ), mul_comm _ (q.succ ! : ℚ), ← mul_assoc, div_mul_eq_mul_div] simp only [add_eq, add_zero, IsUnit.mul_iff, Nat.isUnit_iff, succ.injEq, cast_mul, diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index 8b2189779e182..8d3b0be4e5906 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -478,7 +478,10 @@ theorem c_eq_zero (hz : z ∈ 𝒟ᵒ) (hg : g • z ∈ 𝒟ᵒ) : g 1 0 = 0 := by_contra hc let a := g' 0 0 let d := g' 1 1 - have had : T ^ (-a) * g' = S * T ^ d := by rw [g_eq_of_c_eq_one hc]; group + have had : T ^ (-a) * g' = S * T ^ d := by + rw [g_eq_of_c_eq_one hc] + dsimp [a, d] + group let w := T ^ (-a) • g' • z have h₁ : w = S • T ^ d • z := by simp only [w, ← mul_smul, had] replace h₁ : normSq w < 1 := h₁.symm ▸ normSq_S_smul_lt_one (one_lt_normSq_T_zpow_smul hz d) diff --git a/Mathlib/Order/Fin/Basic.lean b/Mathlib/Order/Fin/Basic.lean index 10983f1527c03..4e47785db09e5 100644 --- a/Mathlib/Order/Fin/Basic.lean +++ b/Mathlib/Order/Fin/Basic.lean @@ -122,7 +122,7 @@ end FromFin /-! #### Monotonicity -/ lemma val_strictMono : StrictMono (val : Fin n → ℕ) := fun _ _ ↦ id -lemma cast_strictMono {k l : ℕ} (h : k = l) : StrictMono (cast h) := fun {_ _} h ↦ h +lemma cast_strictMono {k l : ℕ} (h : k = l) : StrictMono (Fin.cast h) := fun {_ _} h ↦ h lemma strictMono_succ : StrictMono (succ : Fin n → Fin (n + 1)) := fun _ _ ↦ succ_lt_succ lemma strictMono_castLE (h : n ≤ m) : StrictMono (castLE h : Fin n → Fin m) := fun _ _ ↦ id @@ -182,7 +182,7 @@ def orderIsoSubtype : Fin n ≃o {i // i < n} := `castOrderIso eq i` embeds `i` into an equal `Fin` type. -/ @[simps] def castOrderIso (eq : n = m) : Fin n ≃o Fin m where - toEquiv := ⟨cast eq, cast eq.symm, leftInverse_cast eq, rightInverse_cast eq⟩ + toEquiv := ⟨Fin.cast eq, Fin.cast eq.symm, leftInverse_cast eq, rightInverse_cast eq⟩ map_rel_iff' := cast_le_cast eq @[deprecated (since := "2024-05-23")] alias castIso := castOrderIso diff --git a/Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean b/Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean index bf904f8bef341..d19605f9b549b 100644 --- a/Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean +++ b/Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean @@ -326,8 +326,7 @@ theorem natAbs_det_equiv (I : Ideal S) {E : Type*} [EquivLike E S I] [AddEquivCl -- which maps `(S ⧸ I)` to `Π i, ZMod (a i).nat_abs`. haveI : ∀ i, NeZero (a i).natAbs := fun i => ⟨Int.natAbs_ne_zero.mpr (Ideal.smithCoeffs_ne_zero b I hI i)⟩ - simp_rw [Nat.card_congr (Ideal.quotientEquivPiZMod I b hI).toEquiv, Nat.card_pi, - Nat.card_zmod] + simp_rw [Nat.card_congr (Ideal.quotientEquivPiZMod I b hI).toEquiv, Nat.card_pi, Nat.card_zmod, a] /-- Let `b` be a basis for `S` over `ℤ` and `bI` a basis for `I` over `ℤ` of the same dimension. Then an alternative way to compute the norm of `I` is given by taking the determinant of `bI` diff --git a/Mathlib/RingTheory/Kaehler/JacobiZariski.lean b/Mathlib/RingTheory/Kaehler/JacobiZariski.lean index dba3d8955d251..d4c5b8c6ef68f 100644 --- a/Mathlib/RingTheory/Kaehler/JacobiZariski.lean +++ b/Mathlib/RingTheory/Kaehler/JacobiZariski.lean @@ -135,7 +135,7 @@ lemma Cotangent.exact : have hz : z.1 ∈ P.ker.map (Q.toComp P).toAlgHom.toRingHom := e have : Extension.Cotangent.mk (P := (Q.comp P).toExtension) ⟨x, hx'⟩ = Extension.Cotangent.mk z := by - ext; simpa only [comp_vars, val_mk, Ideal.toCotangent_eq, sub_sub_cancel, pow_two] + ext; simpa only [comp_vars, val_mk, Ideal.toCotangent_eq, sub_sub_cancel, pow_two, z] rw [this, ← Submodule.restrictScalars_mem (Q.comp P).Ring, ← Submodule.mem_comap, ← Submodule.span_singleton_le_iff_mem, ← Submodule.map_le_map_iff_of_injective (f := Submodule.subtype _) Subtype.val_injective, Submodule.map_subtype_span_singleton, diff --git a/Mathlib/RingTheory/Polynomial/Chebyshev.lean b/Mathlib/RingTheory/Polynomial/Chebyshev.lean index 24c2e98c66aa7..394b556910dcc 100644 --- a/Mathlib/RingTheory/Polynomial/Chebyshev.lean +++ b/Mathlib/RingTheory/Polynomial/Chebyshev.lean @@ -81,7 +81,7 @@ protected theorem induct (motive : ℤ → Prop) (add_two : ∀ (n : ℕ), motive (↑n + 1) → motive ↑n → motive (↑n + 2)) (neg_add_one : ∀ (n : ℕ), motive (-↑n) → motive (-↑n + 1) → motive (-↑n - 1)) : ∀ (a : ℤ), motive a := - T.induct Unit motive zero one add_two fun n hn hnm => by + T.induct motive zero one add_two fun n hn hnm => by simpa only [Int.negSucc_eq, neg_add] using neg_add_one n hn hnm @[simp] diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 9d8409124777a..9d81bb8ff98ac 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -64,10 +64,11 @@ def omega : ONote := oadd 1 1 0 /-- The ordinal denoted by a notation -/ -@[simp] noncomputable def repr : ONote → Ordinal.{0} | 0 => 0 | oadd e n a => ω ^ repr e * n + repr a +@[simp] theorem repr_zero : repr 0 = 0 := rfl +attribute [simp] repr.eq_1 repr.eq_2 /-- Print `ω^s*n`, omitting `s` if `e = 0` or `e = 1`, and omitting `n` if `n = 1` -/ private def toString_aux (e : ONote) (n : ℕ) (s : String) : String := @@ -112,8 +113,7 @@ instance : WellFoundedRelation ONote := ⟨(· < ·), InvImage.wf repr Ordinal.lt_wf⟩ /-- Convert a `Nat` into an ordinal -/ -@[coe] -def ofNat : ℕ → ONote +@[coe] def ofNat : ℕ → ONote | 0 => 0 | Nat.succ n => oadd 0 n.succPNat 0 @@ -126,17 +126,14 @@ def ofNat : ℕ → ONote @[simp] theorem ofNat_succ (n) : ofNat (Nat.succ n) = oadd 0 n.succPNat 0 := rfl -instance nat (n : ℕ) : OfNat ONote n where +instance (priority := low) nat (n : ℕ) : OfNat ONote n where ofNat := ofNat n -@[simp 1200] -theorem ofNat_one : ofNat 1 = 1 := - rfl +@[simp 1200] theorem ofNat_one : ofNat 1 = 1 := rfl -@[simp] -theorem repr_ofNat (n : ℕ) : repr (ofNat n) = n := by cases n <;> simp +@[simp] theorem repr_ofNat (n : ℕ) : repr (ofNat n) = n := by cases n <;> simp -theorem repr_one : repr (ofNat 1) = (1 : ℕ) := repr_ofNat 1 +@[simp] theorem repr_one : repr 1 = (1 : ℕ) := repr_ofNat 1 theorem omega0_le_oadd (e n a) : ω ^ repr e ≤ repr (oadd e n a) := by refine le_trans ?_ (le_add_right _ _) @@ -174,8 +171,7 @@ theorem eq_of_cmp_eq : ∀ {o₁ o₂}, cmp o₁ o₂ = Ordering.eq → o₁ = o simp protected theorem zero_lt_one : (0 : ONote) < 1 := by - simp only [lt_def, repr, opow_zero, Nat.succPNat_coe, Nat.cast_one, mul_one, add_zero, - zero_lt_one] + simp only [lt_def, repr_zero, repr_one, Nat.cast_one, zero_lt_one] /-- `NFBelow o b` says that `o` is a normal form ordinal notation satisfying `repr o < ω ^ b`. -/ inductive NFBelow : ONote → Ordinal.{0} → Prop @@ -442,7 +438,7 @@ theorem repr_add : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ + o₂) = rep conv at nf => simp [HAdd.hAdd, Add.add] conv in _ + o => simp [HAdd.hAdd, Add.add] cases' h : add a o with e' n' a' <;> - simp only [Add.add, add, addAux, h'.symm, h, add_assoc, repr] at nf h₁ ⊢ + simp only [Add.add, add, addAux, h'.symm, h, add_assoc, repr_zero, repr] at nf h₁ ⊢ have := h₁.fst; haveI := nf.fst; have ee := cmp_compares e e' cases he : cmp e e' <;> simp only [he, Ordering.compares_gt, Ordering.compares_lt, Ordering.compares_eq, repr, gt_iff_lt, PNat.add_coe, Nat.cast_add] at ee ⊢ @@ -568,7 +564,7 @@ theorem repr_mul : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ * o₂) = rep simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos _ omega0_pos).2 (Nat.cast_le.2 n₁.2) by_cases e0 : e₂ = 0 · 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 [Mul.mul, mul, e0, ↓reduceIte, 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] @@ -658,8 +654,7 @@ theorem split_eq_scale_split' : ∀ {o o' m} [NF o], split' o = (o', m) → spli simp only [split_eq_scale_split' h', and_imp] have : 1 + (e - 1) = e := by refine repr_inj.1 ?_ - simp only [repr_add, repr, opow_zero, Nat.succPNat_coe, Nat.cast_one, mul_one, add_zero, - repr_sub] + simp only [repr_add, repr_one, Nat.cast_one, repr_sub] have := mt repr_inj.1 e0 exact Ordinal.add_sub_cancel_of_le <| one_le_iff_ne_zero.2 this intros @@ -684,10 +679,9 @@ theorem nf_repr_split' : ∀ {o o' m} [NF o], split' o = (o', m) → NF o' ∧ r have := mt repr_inj.1 e0 rw [← opow_add, Ordinal.add_sub_cancel_of_le (one_le_iff_ne_zero.2 this)] refine ⟨NF.oadd (by infer_instance) _ ?_, ?_⟩ - · simp at this ⊢ - refine - IH₁.below_of_lt' - ((Ordinal.mul_lt_mul_iff_left omega0_pos).1 <| lt_of_le_of_lt (le_add_right _ m') ?_) + · simp only [opow_one, repr_sub, repr_one, Nat.cast_one] at this ⊢ + refine IH₁.below_of_lt' + ((Ordinal.mul_lt_mul_iff_left omega0_pos).1 <| lt_of_le_of_lt (le_add_right _ m') ?_) rw [← this, ← IH₂] exact h.snd'.repr_lt · rw [this] @@ -717,8 +711,7 @@ theorem nf_repr_split {o o' m} [NF o] (h : split o = (o', m)) : NF o' ∧ repr o cases' nf_repr_split' e with s₁ s₂ rw [split_eq_scale_split' e] at h injection h; substs o' n - simp only [repr_scale, repr, opow_zero, Nat.succPNat_coe, Nat.cast_one, mul_one, add_zero, - opow_one, s₂.symm, and_true] + simp only [repr_scale, repr_one, Nat.cast_one, opow_one, ← s₂, and_true] infer_instance theorem split_dvd {o o' m} [NF o] (h : split o = (o', m)) : ω ∣ repr o' := by @@ -900,7 +893,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₂] + · simp [opowAux2, 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] @@ -924,16 +917,13 @@ theorem repr_opow (o₁ o₂) [NF o₁] [NF o₂] : repr (o₁ ^ o₂) = repr o simp only [opow_def, opow, e₁, r₁, split_eq_scale_split' e₂, opowAux2, repr] cases' k with k · simp [r₂, opow_mul, repr_opow_aux₁ a00 al aa, add_assoc] - · simp? [r₂, opow_add, opow_mul, mul_assoc, add_assoc, -repr, -opow_natCast] says - simp only [mulNat_eq_mul, repr_add, repr_scale, repr_mul, repr_ofNat, opow_add, opow_mul, - mul_assoc, add_assoc, r₂, Nat.cast_add, Nat.cast_one, add_one_eq_succ, opow_succ] - simp only [repr, opow_zero, Nat.succPNat_coe, Nat.cast_one, mul_one, add_zero, opow_one] + · simp [opow, opowAux2, r₂, opow_add, opow_mul, mul_assoc, add_assoc] rw [repr_opow_aux₁ a00 al aa, scale_opowAux] - simp only [repr_mul, repr_scale, repr, opow_zero, Nat.succPNat_coe, Nat.cast_one, mul_one, + simp only [repr_mul, repr_scale, repr, opow_zero, PNat.val_ofNat, Nat.cast_one, mul_one, add_zero, opow_one, opow_mul] rw [← mul_add, ← add_assoc ((ω : Ordinal.{0}) ^ repr a0 * (n : ℕ))] congr 1 - rw [← opow_succ] + rw [← pow_succ, ← opow_natCast, ← opow_natCast] exact (repr_opow_aux₂ _ ad a00 al _ _).2 /-- Given an ordinal, returns: diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index 7938e6809683b..86cf871b98d36 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -68,7 +68,6 @@ import Mathlib.Tactic.DefEqTransformations import Mathlib.Tactic.DeprecateTo import Mathlib.Tactic.DeriveCountable import Mathlib.Tactic.DeriveFintype -import Mathlib.Tactic.DeriveToExpr import Mathlib.Tactic.DeriveTraversable import Mathlib.Tactic.Eqns import Mathlib.Tactic.Eval diff --git a/Mathlib/Tactic/CC/Addition.lean b/Mathlib/Tactic/CC/Addition.lean index 6e7a9fe150ad2..4ad4dc22d38f0 100644 --- a/Mathlib/Tactic/CC/Addition.lean +++ b/Mathlib/Tactic/CC/Addition.lean @@ -398,8 +398,8 @@ equality `r*a = s*b`. -/ def superposeAC (ts a : ACApps) (tsEqa : DelayedExpr) : CCM Unit := do let .apps op args := ts | return for hi : i in [:args.size] do - if i == 0 || (args[i]'hi.2) != (args[i - 1]'(Nat.lt_of_le_of_lt (i.sub_le 1) hi.2)) then - let some ent := (← get).acEntries.find? (args[i]'hi.2) | failure + if i == 0 || args[i] != (args[i - 1]'(Nat.lt_of_le_of_lt (i.sub_le 1) hi.2.1)) then + let some ent := (← get).acEntries.find? args[i] | failure let occs := ent.RLHSOccs for tr in occs do let .apps optr _ := tr | continue @@ -607,7 +607,7 @@ partial def internalizeAppLit (e : Expr) : CCM Unit := do pinfo := (← getFunInfoNArgs fn apps.size).paramInfo.toList if state.hoFns.isSome && fn.isConst && !(state.hoFns.iget.contains fn.constName) then for h : i in [:apps.size] do - let arg := (apps[i]'h.2).appArg! + let arg := apps[i].appArg! addOccurrence e arg false if pinfo.head?.any ParamInfo.isInstImplicit then -- We do not recurse on instances when `(← get).config.ignoreInstances` is `true`. @@ -625,13 +625,13 @@ partial def internalizeAppLit (e : Expr) : CCM Unit := do -- Expensive case where we store a quadratic number of occurrences, -- as described in the paper "Congruence Closure in Internsional Type Theory" for h : i in [:apps.size] do - let curr := apps[i]'h.2 + let curr := apps[i] let .app currFn currArg := curr | unreachable! if i < apps.size - 1 then mkEntry curr false for h : j in [i:apps.size] do - addOccurrence (apps[j]'h.2) currArg false - addOccurrence (apps[j]'h.2) currFn false + addOccurrence apps[j] currArg false + addOccurrence apps[j] currFn false if pinfo.head?.any ParamInfo.isInstImplicit then -- We do not recurse on instances when `(← get).config.ignoreInstances` is `true`. mkEntry currArg false diff --git a/Mathlib/Tactic/CC/Datatypes.lean b/Mathlib/Tactic/CC/Datatypes.lean index d45f190596b92..6dc3e10fd6aa5 100644 --- a/Mathlib/Tactic/CC/Datatypes.lean +++ b/Mathlib/Tactic/CC/Datatypes.lean @@ -156,7 +156,7 @@ scoped instance : Ord ACApps where compare op₁ op₂ |>.then <| compare args₁.size args₂.size |>.dthen fun hs => Id.run do have hs := Batteries.BEqCmp.cmp_iff_eq.mp hs for hi : i in [:args₁.size] do - have hi := hi.right; let o := compare (args₁[i]'hi) (args₂[i]'(hs ▸ hi)) + have hi := hi.right; let o := compare args₁[i] (args₂[i]'(hs ▸ hi.1)) if o != .eq then return o return .eq @@ -583,10 +583,10 @@ def getVarWithLeastOccs (ccs : CCState) (e : ACApps) (inLHS : Bool) : Option Exp let mut r := args[0]? let mut numOccs := r.casesOn 0 fun r' => ccs.getNumROccs r' inLHS for hi : i in [1:args.size] do - if (args[i]'hi.2) != (args[i - 1]'(Nat.lt_of_le_of_lt (i.sub_le 1) hi.2)) then - let currOccs := ccs.getNumROccs (args[i]'hi.2) inLHS + if args[i] != (args[i - 1]'(Nat.lt_of_le_of_lt (i.sub_le 1) hi.2.1)) then + let currOccs := ccs.getNumROccs args[i] inLHS if currOccs < numOccs then - r := (args[i]'hi.2) + r := args[i] numOccs := currOccs return r | .ofExpr e => e diff --git a/Mathlib/Tactic/CC/MkProof.lean b/Mathlib/Tactic/CC/MkProof.lean index c3827b0526b16..5f17a75846cc5 100644 --- a/Mathlib/Tactic/CC/MkProof.lean +++ b/Mathlib/Tactic/CC/MkProof.lean @@ -73,7 +73,7 @@ partial def isCongruent (e₁ e₂ : Expr) : CCM Bool := do e₂.withApp fun f₂ args₂ => do if ha : args₁.size = args₂.size then for hi : i in [:args₁.size] do - if (← getRoot (args₁[i]'hi.2)) != (← getRoot (args₂[i]'(ha.symm ▸ hi.2))) then + if (← getRoot args₁[i]) != (← getRoot (args₂[i]'(ha.symm ▸ hi.2.1))) then return false if f₁ == f₂ then return true else if (← getRoot f₁) != (← getRoot f₂) then @@ -289,13 +289,13 @@ partial def mkCongrProofCore (lhs rhs : Expr) (heqProofs : Bool) : CCM Expr := d let mut lemmaArgs : Array Expr := #[] for hi : i in [:lhsArgs.size] do guard !kindsIt.isEmpty - lemmaArgs := lemmaArgs.push (lhsArgs[i]'hi.2) |>.push (rhsArgs[i]'(ha.symm ▸ hi.2)) + lemmaArgs := lemmaArgs.push lhsArgs[i] |>.push (rhsArgs[i]'(ha.symm ▸ hi.2.1)) if kindsIt[0]! matches CongrArgKind.heq then - let some p ← getHEqProof (lhsArgs[i]'hi.2) (rhsArgs[i]'(ha.symm ▸ hi.2)) | failure + let some p ← getHEqProof lhsArgs[i] (rhsArgs[i]'(ha.symm ▸ hi.2.1)) | failure lemmaArgs := lemmaArgs.push p else guard (kindsIt[0]! matches .eq) - let some p ← getEqProof (lhsArgs[i]'hi.2) (rhsArgs[i]'(ha.symm ▸ hi.2)) | failure + let some p ← getEqProof lhsArgs[i] (rhsArgs[i]'(ha.symm ▸ hi.2.1)) | failure lemmaArgs := lemmaArgs.push p kindsIt := kindsIt.eraseIdx! 0 let mut r := mkAppN specLemma.proof lemmaArgs @@ -606,8 +606,8 @@ expression. -/ def simplifyACStep (e : ACApps) : CCM (Option (ACApps × DelayedExpr)) := do if let .apps _ args := e then for h : i in [:args.size] do - if i == 0 || (args[i]'h.2) != (args[i - 1]'(Nat.lt_of_le_of_lt (i.sub_le 1) h.2)) then - let some ae := (← get).acEntries.find? (args[i]'h.2) | failure + if i == 0 || args[i] != (args[i - 1]'(Nat.lt_of_le_of_lt (i.sub_le 1) h.2.1)) then + let some ae := (← get).acEntries.find? args[i] | failure let occs := ae.RLHSOccs let mut Rlhs? : Option ACApps := none for Rlhs in occs do diff --git a/Mathlib/Tactic/Common.lean b/Mathlib/Tactic/Common.lean index b59d768f32d0c..00ceda6c80a6d 100644 --- a/Mathlib/Tactic/Common.lean +++ b/Mathlib/Tactic/Common.lean @@ -13,7 +13,6 @@ import Plausible import ImportGraph.Imports -- Import common Batteries tactics and commands -import Batteries.Tactic.Where import Batteries.Tactic.Basic import Batteries.Tactic.HelpCmd @@ -47,7 +46,6 @@ import Mathlib.Tactic.Conv import Mathlib.Tactic.Convert import Mathlib.Tactic.DefEqTransformations import Mathlib.Tactic.DeprecateTo -import Mathlib.Tactic.DeriveToExpr import Mathlib.Tactic.Eqns import Mathlib.Tactic.ExistsI import Mathlib.Tactic.ExtractGoal diff --git a/Mathlib/Tactic/DeriveToExpr.lean b/Mathlib/Tactic/DeriveToExpr.lean deleted file mode 100644 index d82a0df118816..0000000000000 --- a/Mathlib/Tactic/DeriveToExpr.lean +++ /dev/null @@ -1,230 +0,0 @@ -/- -Copyright (c) 2023 Kyle Miller. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kyle Miller --/ -import Lean.Elab.Deriving.Ord -import Mathlib.Tactic.ToLevel - -/-! -# A `ToExpr` derive handler - -This module defines a `ToExpr` derive handler for inductive types. It supports mutually inductive -types as well. - -The `ToExpr` derive handlers support universe level polymorphism. This is implemented using the -`Lean.ToLevel` class. To use `ToExpr` in places where there is universe polymorphism, make sure -to have a `[ToLevel.{u}]` instance available. - -**Warning:** Import `Mathlib.Tactic.ToExpr` instead of this one. This ensures that you are using -the universe polymorphic `ToExpr` instances that override the ones from Lean 4 core. - -Implementation note: this derive handler was originally modeled after the `Repr` derive handler. --/ - -namespace Mathlib.Deriving.ToExpr - -open Lean Elab Lean.Parser.Term -open Meta Command Deriving - -/-- Specialization of `Lean.Elab.Deriving.mkHeader` for `ToExpr`. -/ -def mkToExprHeader (indVal : InductiveVal) : TermElabM Header := do - -- The auxiliary functions we produce are `indtype -> Expr`. - let header ← mkHeader ``ToExpr 1 indVal - return header - -/-- Give a term that is equivalent to `(term| mkAppN $f #[$args,*])`. -As an optimization, `mkAppN` is pre-expanded out to use `Expr.app` directly. -/ -def mkAppNTerm (f : Term) (args : Array Term) : MetaM Term := - args.foldlM (fun a b => `(Expr.app $a $b)) f - -/-- Create the body of the `toExpr` function -for the `ToExpr` instance, which is a `match` expression -that calls `toExpr` and `toTypeExpr` to assemble an expression for a given term. -For recursive inductive types, `auxFunName` refers to the `ToExpr` instance -for the current type. -For mutually recursive types, we rely on the local instances set up by `mkLocalInstanceLetDecls`. -/ -def mkToExprBody (header : Header) (indVal : InductiveVal) (auxFunName : Name) : - TermElabM Term := do - let discrs ← mkDiscrs header indVal - let alts ← mkAlts - `(match $[$discrs],* with $alts:matchAlt*) -where - /-- Create the `match` cases, one per constructor. -/ - mkAlts : TermElabM (Array (TSyntax ``matchAlt)) := do - let mut alts := #[] - for ctorName in indVal.ctors do - let ctorInfo ← getConstInfoCtor ctorName - let alt ← forallTelescopeReducing ctorInfo.type fun xs _ => do - let mut patterns := #[] - -- add `_` pattern for indices - for _ in [:indVal.numIndices] do - patterns := patterns.push (← `(_)) - let mut ctorArgs := #[] - let mut rhsArgs : Array Term := #[] - let mkArg (x : Expr) (a : Term) : TermElabM Term := do - if (← inferType x).isAppOf indVal.name then - `($(mkIdent auxFunName) $a) - else if ← Meta.isType x then - `(toTypeExpr $a) - else - `(toExpr $a) - -- add `_` pattern for inductive parameters, which are inaccessible - for i in [:ctorInfo.numParams] do - let a := mkIdent header.argNames[i]! - ctorArgs := ctorArgs.push (← `(_)) - rhsArgs := rhsArgs.push <| ← mkArg xs[i]! a - for i in [:ctorInfo.numFields] do - let a := mkIdent (← mkFreshUserName `a) - ctorArgs := ctorArgs.push a - rhsArgs := rhsArgs.push <| ← mkArg xs[ctorInfo.numParams + i]! a - patterns := patterns.push (← `(@$(mkIdent ctorName):ident $ctorArgs:term*)) - let levels ← indVal.levelParams.toArray.mapM (fun u => `(toLevel.{$(mkIdent u)})) - let rhs : Term ← - mkAppNTerm (← `(Expr.const $(quote ctorInfo.name) [$levels,*])) rhsArgs - `(matchAltExpr| | $[$patterns:term],* => $rhs) - alts := alts.push alt - return alts - -/-- Create the body of the `toTypeExpr` function for the `ToExpr` instance. -Calls `toExpr` and `toTypeExpr` to the arguments to the type constructor. -/ -def mkToTypeExpr (argNames : Array Name) (indVal : InductiveVal) : TermElabM Term := do - let levels ← indVal.levelParams.toArray.mapM (fun u => `(toLevel.{$(mkIdent u)})) - forallTelescopeReducing indVal.type fun xs _ => do - let mut args : Array Term := #[] - for i in [:xs.size] do - let x := xs[i]! - let a := mkIdent argNames[i]! - if ← Meta.isType x then - args := args.push <| ← `(toTypeExpr $a) - else - args := args.push <| ← `(toExpr $a) - mkAppNTerm (← `((Expr.const $(quote indVal.name) [$levels,*]))) args - -/-- -For mutually recursive inductive types, the strategy is to have local `ToExpr` instances in scope -for each of the inductives when defining each instance. -This way, each instance can freely use `toExpr` and `toTypeExpr` for each of the other types. - -Note that each instance gets its own definition of each of the others' `toTypeExpr` fields. -(This is working around the fact that the `Deriving.Context` API assumes -that each instance in mutual recursion only has a single auxiliary definition. -There are other ways to work around it, but `toTypeExpr` implementations -are very simple, so duplicating them seemed to be OK.) -/ -def mkLocalInstanceLetDecls (ctx : Deriving.Context) (argNames : Array Name) : - TermElabM (Array (TSyntax ``Parser.Term.letDecl)) := do - let mut letDecls := #[] - for i in [:ctx.typeInfos.size] do - let indVal := ctx.typeInfos[i]! - let auxFunName := ctx.auxFunNames[i]! - let currArgNames ← mkInductArgNames indVal - let numParams := indVal.numParams - let currIndices := currArgNames[numParams:] - let binders ← mkImplicitBinders currIndices - let argNamesNew := argNames[:numParams] ++ currIndices - let indType ← mkInductiveApp indVal argNamesNew - let instName ← mkFreshUserName `localinst - let toTypeExpr ← mkToTypeExpr argNames indVal - let letDecl ← `(Parser.Term.letDecl| $(mkIdent instName):ident $binders:implicitBinder* : - ToExpr $indType := - { toExpr := $(mkIdent auxFunName), toTypeExpr := $toTypeExpr }) - letDecls := letDecls.push letDecl - return letDecls - -/-- Fix the output of `mkInductiveApp` to explicitly reference universe levels. -/ -def fixIndType (indVal : InductiveVal) (t : Term) : TermElabM Term := - match t with - | `(@$f $args*) => - let levels := indVal.levelParams.toArray.map mkIdent - `(@$f.{$levels,*} $args*) - | _ => throwError "(internal error) expecting output of `mkInductiveApp`" - -/-- Make `ToLevel` instance binders for all the level variables. -/ -def mkToLevelBinders (indVal : InductiveVal) : TermElabM (TSyntaxArray ``instBinderF) := do - indVal.levelParams.toArray.mapM (fun u => `(instBinderF| [ToLevel.{$(mkIdent u)}])) - -open TSyntax.Compat in -/-- Make a `toExpr` function for the given inductive type. -The implementations of each `toExpr` function for a (mutual) inductive type -are given as top-level private definitions. -These end up being assembled into `ToExpr` instances in `mkInstanceCmds`. -For mutual inductive types, -then each of the other types' `ToExpr` instances are provided as local instances, -to wire together the recursion (this necessitates these auxiliary definitions being `partial`). -/ -def mkAuxFunction (ctx : Deriving.Context) (i : Nat) : TermElabM Command := do - let auxFunName := ctx.auxFunNames[i]! - let indVal := ctx.typeInfos[i]! - let header ← mkToExprHeader indVal - let mut body ← mkToExprBody header indVal auxFunName - if ctx.usePartial then - let letDecls ← mkLocalInstanceLetDecls ctx header.argNames - body ← mkLet letDecls body - -- We need to alter the last binder (the one for the "target") to have explicit universe levels - -- so that the `ToLevel` instance arguments can use them. - let addLevels binder := - match binder with - | `(bracketedBinderF| ($a : $ty)) => do `(bracketedBinderF| ($a : $(← fixIndType indVal ty))) - | _ => throwError "(internal error) expecting inst binder" - let binders := header.binders.pop - ++ (← mkToLevelBinders indVal) - ++ #[← addLevels header.binders.back!] - let levels := indVal.levelParams.toArray.map mkIdent - if ctx.usePartial then - `(private partial def $(mkIdent auxFunName):ident.{$levels,*} $binders:bracketedBinder* : - Expr := $body:term) - else - `(private def $(mkIdent auxFunName):ident.{$levels,*} $binders:bracketedBinder* : - Expr := $body:term) - -/-- Create all the auxiliary functions using `mkAuxFunction` for the (mutual) inductive type(s). -Wraps the resulting definition commands in `mutual ... end`. -/ -def mkMutualBlock (ctx : Deriving.Context) : TermElabM Syntax := do - let mut auxDefs := #[] - for i in [:ctx.typeInfos.size] do - auxDefs := auxDefs.push (← mkAuxFunction ctx i) - `(mutual $auxDefs:command* end) - -open TSyntax.Compat in -/-- Assuming all of the auxiliary definitions exist, create all the `instance` commands -for the `ToExpr` instances for the (mutual) inductive type(s). -/ -def mkInstanceCmds (ctx : Deriving.Context) (typeNames : Array Name) : - TermElabM (Array Command) := do - let mut instances := #[] - for i in [:ctx.typeInfos.size] do - let indVal := ctx.typeInfos[i]! - if typeNames.contains indVal.name then - let auxFunName := ctx.auxFunNames[i]! - let argNames ← mkInductArgNames indVal - let binders ← mkImplicitBinders argNames - let binders := binders ++ (← mkInstImplicitBinders ``ToExpr indVal argNames) - let binders := binders ++ (← mkToLevelBinders indVal) - let indType ← fixIndType indVal (← mkInductiveApp indVal argNames) - let toTypeExpr ← mkToTypeExpr argNames indVal - let levels := indVal.levelParams.toArray.map mkIdent - let instCmd ← `(instance $binders:implicitBinder* : ToExpr $indType where - toExpr := $(mkIdent auxFunName).{$levels,*} - toTypeExpr := $toTypeExpr) - instances := instances.push instCmd - return instances - -/-- Returns all the commands generated by `mkMutualBlock` and `mkInstanceCmds`. -/ -def mkToExprInstanceCmds (declNames : Array Name) : TermElabM (Array Syntax) := do - let ctx ← mkContext "toExpr" declNames[0]! - let cmds := #[← mkMutualBlock ctx] ++ (← mkInstanceCmds ctx declNames) - trace[Elab.Deriving.toExpr] "\n{cmds}" - return cmds - -/-- The main entry point to the `ToExpr` derive handler. -/ -def mkToExprInstanceHandler (declNames : Array Name) : CommandElabM Bool := do - if (← declNames.allM isInductive) && declNames.size > 0 then - let cmds ← liftTermElabM <| mkToExprInstanceCmds declNames - cmds.forM elabCommand - return true - else - return false - -initialize - registerDerivingHandler ``Lean.ToExpr mkToExprInstanceHandler - registerTraceClass `Elab.Deriving.toExpr - -end Mathlib.Deriving.ToExpr diff --git a/Mathlib/Tactic/ITauto.lean b/Mathlib/Tactic/ITauto.lean index f87cbcddd6ec3..f9e460c2a588f 100644 --- a/Mathlib/Tactic/ITauto.lean +++ b/Mathlib/Tactic/ITauto.lean @@ -6,7 +6,6 @@ Authors: Mario Carneiro import Batteries.Tactic.Exact import Batteries.Tactic.Init import Mathlib.Logic.Basic -import Mathlib.Tactic.DeriveToExpr import Mathlib.Util.AtomM import Qq diff --git a/Mathlib/Tactic/Linarith/Frontend.lean b/Mathlib/Tactic/Linarith/Frontend.lean index a63112c459aff..e76aaf4acb162 100644 --- a/Mathlib/Tactic/Linarith/Frontend.lean +++ b/Mathlib/Tactic/Linarith/Frontend.lean @@ -223,7 +223,7 @@ abbrev ExprMultiMap α := Array (Expr × List α) (If the key is not in the map it returns `self.size` as the index.) -/ def ExprMultiMap.find {α : Type} (self : ExprMultiMap α) (k : Expr) : MetaM (Nat × List α) := do for h : i in [:self.size] do - let (k', vs) := self[i]'h.2 + let (k', vs) := self[i] if ← isDefEq k' k then return (i, vs) return (self.size, []) @@ -233,7 +233,7 @@ in the map. -/ def ExprMultiMap.insert {α : Type} (self : ExprMultiMap α) (k : Expr) (v : α) : MetaM (ExprMultiMap α) := do for h : i in [:self.size] do - if ← isDefEq (self[i]'h.2).1 k then + if ← isDefEq self[i].1 k then return self.modify i fun (k, vs) => (k, v::vs) return self.push (k, [v]) diff --git a/Mathlib/Tactic/Linter/FlexibleLinter.lean b/Mathlib/Tactic/Linter/FlexibleLinter.lean index f4ad7305528d7..f2a568cca5b34 100644 --- a/Mathlib/Tactic/Linter/FlexibleLinter.lean +++ b/Mathlib/Tactic/Linter/FlexibleLinter.lean @@ -295,7 +295,7 @@ def flexible : Std.HashSet Name := `Mathlib.Tactic.normNum, `linarith, `nlinarith, - ``Lean.Parser.Tactic.tacticNorm_cast_, + ``Lean.Parser.Tactic.tacticNorm_cast__, `Aesop.Frontend.Parser.aesopTactic, `Mathlib.Tactic.Tauto.tauto, `Mathlib.Meta.FunProp.funPropTacStx, diff --git a/Mathlib/Tactic/Linter/HaveLetLinter.lean b/Mathlib/Tactic/Linter/HaveLetLinter.lean index b88a5344a4265..299b1766d115f 100644 --- a/Mathlib/Tactic/Linter/HaveLetLinter.lean +++ b/Mathlib/Tactic/Linter/HaveLetLinter.lean @@ -49,10 +49,9 @@ register_option linter.haveLet : Nat := { namespace haveLet /-- find the `have` syntax. -/ -partial def isHave? : Syntax → Bool | .node _ ``Lean.Parser.Tactic.tacticHave_ _ => true - |_ => false + | _ => false end haveLet diff --git a/Mathlib/Tactic/SuccessIfFailWithMsg.lean b/Mathlib/Tactic/SuccessIfFailWithMsg.lean index 3881cd09cfd2c..b3ef1a312fae9 100644 --- a/Mathlib/Tactic/SuccessIfFailWithMsg.lean +++ b/Mathlib/Tactic/SuccessIfFailWithMsg.lean @@ -29,7 +29,7 @@ syntax (name := successIfFailWithMsg) "success_if_fail_with_msg " term:max tacti /-- Evaluates `tacs` and succeeds only if `tacs` both fails and throws an error equal (as a string) to `msg`. -/ def successIfFailWithMessage {s α : Type} {m : Type → Type} - [Monad m] [MonadLiftT IO m] [MonadBacktrack s m] [MonadError m] + [Monad m] [MonadLiftT BaseIO m] [MonadBacktrack s m] [MonadError m] (msg : String) (tacs : m α) (ref : Option Syntax := none) : m Unit := do let s ← saveState let err ← diff --git a/Mathlib/Tactic/ToExpr.lean b/Mathlib/Tactic/ToExpr.lean index 8997800729532..72f4fd3b7ee87 100644 --- a/Mathlib/Tactic/ToExpr.lean +++ b/Mathlib/Tactic/ToExpr.lean @@ -3,50 +3,12 @@ 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.Tactic.DeriveToExpr import Mathlib.Util.WhatsNew -/-! # `ToExpr` instances for Mathlib - -This module should be imported by any module that intends to define `ToExpr` instances. -It provides necessary dependencies (the `Lean.ToLevel` class) and it also overrides the instances -that come from core Lean 4 that do not handle universe polymorphism. -(See the module `Lean.ToExpr` for the instances that are overridden.) - -In addition, we provide some additional `ToExpr` instances for core definitions. +/-! +# `ToExpr` instances for Mathlib -/ -section override -- Note: this section uses `autoImplicit` pervasively -namespace Lean - -attribute [-instance] Lean.instToExprOption - -set_option autoImplicit true in -deriving instance ToExpr for Option - -attribute [-instance] Lean.instToExprList - -set_option autoImplicit true in -deriving instance ToExpr for List - -attribute [-instance] Lean.instToExprArray - -universe u in -instance {α : Type u} [ToExpr α] [ToLevel.{u}] : ToExpr (Array α) := - let type := toTypeExpr α - { toExpr := fun as => mkApp2 (mkConst ``List.toArray [toLevel.{u}]) type (toExpr as.toList) - toTypeExpr := mkApp (mkConst ``Array [toLevel.{u}]) type } - -attribute [-instance] Lean.instToExprProd - -set_option autoImplicit true in -deriving instance ToExpr for Prod - -deriving instance ToExpr for System.FilePath - -end Lean -end override - namespace Mathlib open Lean @@ -62,7 +24,6 @@ instance [ToLevel.{u}] : ToExpr PUnit.{u+1} where deriving instance ToExpr for String.Pos deriving instance ToExpr for Substring deriving instance ToExpr for SourceInfo -deriving instance ToExpr for Syntax.Preresolved deriving instance ToExpr for Syntax open DataValue in @@ -86,12 +47,10 @@ instance : ToExpr MData where toExpr := toExprMData toTypeExpr := mkConst ``MData -deriving instance ToExpr for FVarId deriving instance ToExpr for MVarId deriving instance ToExpr for LevelMVarId deriving instance ToExpr for Level deriving instance ToExpr for BinderInfo -deriving instance ToExpr for Literal deriving instance ToExpr for Expr end Mathlib diff --git a/Mathlib/Tactic/ToLevel.lean b/Mathlib/Tactic/ToLevel.lean index 17b243651eb54..fd85a5198e967 100644 --- a/Mathlib/Tactic/ToLevel.lean +++ b/Mathlib/Tactic/ToLevel.lean @@ -17,31 +17,6 @@ override the ones from Lean 4 core. namespace Lean -/-- A class to create `Level` expressions that denote particular universe levels in Lean. -`Lean.ToLevel.toLevel.{u}` evaluates to a `Lean.Level` term representing `u` -/ -@[pp_with_univ] -class ToLevel.{u} where - /-- A `Level` that represents the universe level `u`. -/ - toLevel : Level - /-- The universe itself. This is only here to avoid the "unused universe parameter" error. -/ - univ : Type u := Sort u -export ToLevel (toLevel) attribute [pp_with_univ] toLevel -instance : ToLevel.{0} where - toLevel := .zero - -universe u v - -instance [ToLevel.{u}] : ToLevel.{u+1} where - toLevel := .succ toLevel.{u} - -/-- `ToLevel` for `max u v`. This is not an instance since it causes divergence. -/ -def ToLevel.max [ToLevel.{u}] [ToLevel.{v}] : ToLevel.{max u v} where - toLevel := .max toLevel.{u} toLevel.{v} - -/-- `ToLevel` for `imax u v`. This is not an instance since it causes divergence. -/ -def ToLevel.imax [ToLevel.{u}] [ToLevel.{v}] : ToLevel.{imax u v} where - toLevel := .imax toLevel.{u} toLevel.{v} - end Lean diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index 2af0c9f25bf29..fc08069c2269f 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -653,8 +653,7 @@ theorem GoodProducts.spanFin [WellFoundedLT I] : | cons b bs => apply le_of_lt rw [List.chain'_cons] at ha - have hlex := List.lt.head bs (b :: bs) ha.1 - exact (List.lt_iff_lex_lt _ _).mp hlex + exact (List.lt_iff_lex_lt _ _).mp (List.Lex.rel ha.1) end Fin @@ -783,7 +782,7 @@ def Products.nil : Products I := ⟨[], by simp only [List.chain'_nil]⟩ 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 + simp only [Set.mem_setOf_eq, lt_iff_lex_lt, nil, List.not_lex_nil] at h instance {α : Type*} [TopologicalSpace α] [Nonempty α] : Nontrivial (LocallyConstant α ℤ) := ⟨0, 1, ne_of_apply_ne DFunLike.coe <| (Function.const_injective (β := ℤ)).ne zero_ne_one⟩ @@ -1588,7 +1587,7 @@ theorem chain'_cons_of_lt (l : MaxProducts C ho) refine lt_of_le_of_lt (Products.head!_le_of_lt hq (q.val.ne_nil_of_mem ha)) ?_ by_cases hM : l.val.Tail.val = [] · rw [Products.lt_iff_lex_lt, hM] at hq - simp only [List.Lex.not_nil_right] at hq + simp only [List.not_lex_nil] at hq · have := l.val.prop 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) diff --git a/MathlibTest/DeriveToExpr.lean b/MathlibTest/DeriveToExpr.lean index 54e4bae668ad1..1fa2cb66f0338 100644 --- a/MathlibTest/DeriveToExpr.lean +++ b/MathlibTest/DeriveToExpr.lean @@ -1,4 +1,4 @@ -import Mathlib.Tactic.DeriveToExpr +import Lean namespace DeriveToExprTests open Lean @@ -57,7 +57,6 @@ instance {α : Type u} [ToExpr α] [ToLevel.{u+1}] : ToExpr (Bool → α) where deriving instance ToExpr for Bar -set_option linter.unusedTactic false in example : True := by run_tac do let f : Bool → Nat | false => 0 | true => 1 diff --git a/MathlibTest/HashCommandLinter.lean b/MathlibTest/HashCommandLinter.lean index 4193342406c14..cf5304e11e734 100644 --- a/MathlibTest/HashCommandLinter.lean +++ b/MathlibTest/HashCommandLinter.lean @@ -2,6 +2,12 @@ import Lean.Elab.GuardMsgs import Mathlib.Tactic.AdaptationNote import Mathlib.Tactic.Linter.HashCommandLinter + +-- #adaptation_note +-- The hashCommand linter started failing after https://github.com/leanprover/lean4/pull/6299 +-- Disabling aync elaboration fixes it, but of course we're not going to do that globally. +set_option Elab.async false + set_option linter.hashCommand true section ignored_commands diff --git a/MathlibTest/HaveLetLinter.lean b/MathlibTest/HaveLetLinter.lean index ec7409b66ccde..96d2aae3e74e1 100644 --- a/MathlibTest/HaveLetLinter.lean +++ b/MathlibTest/HaveLetLinter.lean @@ -1,6 +1,12 @@ import Mathlib.Tactic.Linter.HaveLetLinter import Mathlib.Tactic.Tauto + +-- #adaptation_note +-- The haveLet linter started failing after https://github.com/leanprover/lean4/pull/6299 +-- Disabling aync elaboration fixes it, but of course we're not going to do that globally. +set_option Elab.async false + set_option linter.haveLet 1 /-- diff --git a/MathlibTest/TransImports.lean b/MathlibTest/TransImports.lean index 964187009033d..c2cc31e073e1f 100644 --- a/MathlibTest/TransImports.lean +++ b/MathlibTest/TransImports.lean @@ -1,10 +1,10 @@ import Mathlib.Util.TransImports /-- -info: 'MathlibTest.TransImports' has at most 500 transitive imports +info: 'MathlibTest.TransImports' has at most 1000 transitive imports 3 starting with "Lean.Elab.I": [Lean.Elab.InfoTree, Lean.Elab.InfoTree.Main, Lean.Elab.InfoTree.Types] -/ #guard_msgs in -#trans_imports "Lean.Elab.I" at_most 500 +#trans_imports "Lean.Elab.I" at_most 1000 diff --git a/MathlibTest/toAdditive.lean b/MathlibTest/toAdditive.lean index a1f60be919017..6e9807fa7979b 100644 --- a/MathlibTest/toAdditive.lean +++ b/MathlibTest/toAdditive.lean @@ -424,7 +424,7 @@ run_cmd do unless findTranslation? (← getEnv) `localize.s == some `add_localize.s do throwError "3" /-- -warning: The source declaration one_eq_one was given the simp-attribute(s) reduce_mod_char, simp before calling @[to_additive]. The preferred method is to use something like `@[to_additive (attr := reduce_mod_char, simp)]` to apply the attribute to both one_eq_one and the target declaration zero_eq_zero. +warning: The source declaration one_eq_one was given the simp-attribute(s) simp, reduce_mod_char before calling @[to_additive]. The preferred method is to use something like `@[to_additive (attr := simp, reduce_mod_char)]` to apply the attribute to both one_eq_one and the target declaration zero_eq_zero. note: this linter can be disabled with `set_option linter.existingAttributeWarning false` -/ #guard_msgs in diff --git a/lake-manifest.json b/lake-manifest.json index b6975fa1f851e..8a7f0e3ba49a2 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,10 +5,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "2c57364ef83406ea86d0f78ce3e342079a2fece5", + "rev": "1622a8693b31523c8f82db48e01b14c74bc1f155", "name": "plausible", "manifestFile": "lake-manifest.json", - "inputRev": "v4.15.0", + "inputRev": "v4.16.0-rc1", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/LeanSearchClient", @@ -25,30 +25,30 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "9a0b533c2fbd6195df067630be18e11e4349051c", + "rev": "9cb79405471ae931ac718231d6299bfaffef9087", "name": "importGraph", "manifestFile": "lake-manifest.json", - "inputRev": "v4.15.0", + "inputRev": "v4.16.0-rc1", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/ProofWidgets4", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "2b000e02d50394af68cfb4770a291113d94801b5", + "rev": "07f60e90998dfd6592688a14cd67bd4e384b77b2", "name": "proofwidgets", "manifestFile": "lake-manifest.json", - "inputRev": "v0.0.48", + "inputRev": "v0.0.50", "inherited": false, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover-community/aesop", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "2689851f387bb2cef351e6825fe94a56a304ca13", + "rev": "79402ad9ab4be9a2286701a9880697e2351e4955", "name": "aesop", "manifestFile": "lake-manifest.json", - "inputRev": "v4.15.0", + "inputRev": "v4.16.0-rc1", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/quote4", @@ -65,10 +65,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "e8dc5fc16c625fc4fe08f42d625523275ddbbb4b", + "rev": "8ce422eb59adf557fac184f8b1678c75fa03075c", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "v4.15.0", + "inputRev": "v4.16.0-rc1", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lakefile.lean b/lakefile.lean index 44bbb99b71251..a9a535eb971a2 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -7,13 +7,13 @@ open Lake DSL ## Mathlib dependencies on upstream projects -/ -require "leanprover-community" / "batteries" @ git "v4.15.0" +require "leanprover-community" / "batteries" @ git "v4.16.0-rc1" require "leanprover-community" / "Qq" @ git "v4.15.0" -require "leanprover-community" / "aesop" @ git "v4.15.0" -require "leanprover-community" / "proofwidgets" @ git "v0.0.48" -require "leanprover-community" / "importGraph" @ git "v4.15.0" +require "leanprover-community" / "aesop" @ git "v4.16.0-rc1" +require "leanprover-community" / "proofwidgets" @ git "v0.0.50" +require "leanprover-community" / "importGraph" @ git "v4.16.0-rc1" require "leanprover-community" / "LeanSearchClient" @ git "main" -require "leanprover-community" / "plausible" @ git "v4.15.0" +require "leanprover-community" / "plausible" @ git "v4.16.0-rc1" /-! ## Options for building mathlib diff --git a/lean-toolchain b/lean-toolchain index d0eb99ff68d76..62ccd7170eb94 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.15.0 +leanprover/lean4:v4.16.0-rc1 diff --git a/scripts/noshake.json b/scripts/noshake.json index 784455d70e40a..8a935b90e4553 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -90,7 +90,6 @@ "Mathlib.Tactic.Core", "Mathlib.Tactic.DefEqTransformations", "Mathlib.Tactic.DeriveFintype", - "Mathlib.Tactic.DeriveToExpr", "Mathlib.Tactic.Eqns", "Mathlib.Tactic.ExistsI", "Mathlib.Tactic.ExtendDoc", @@ -275,7 +274,6 @@ "Mathlib.Tactic.Positivity.Core"], "Mathlib.Tactic.FinCases": ["Mathlib.Data.Fintype.Basic"], "Mathlib.Tactic.Eval": ["Qq.Macro"], - "Mathlib.Tactic.DeriveToExpr": ["Lean.Elab.Deriving.Ord"], "Mathlib.Tactic.DeriveFintype": ["Mathlib.Data.Fintype.Basic", "Mathlib.Data.Fintype.Sigma", @@ -384,6 +382,8 @@ "Mathlib.Data.List.Range": ["Batteries.Data.Nat.Lemmas"], "Mathlib.Data.List.ProdSigma": ["Batteries.Data.Nat.Lemmas", "Mathlib.Data.List.Basic"], + "Mathlib.Data.List.Permutation": + ["Mathlib.Data.List.Flatten", "Mathlib.Data.List.Lemmas"], "Mathlib.Data.List.Lemmas": ["Mathlib.Data.List.InsertIdx"], "Mathlib.Data.List.Defs": ["Batteries.Data.RBMap.Basic"], "Mathlib.Data.List.Basic": From 3c125512032771407147a2b35348eb1f4c3a12ba Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 5 Jan 2025 04:19:38 +0000 Subject: [PATCH 151/189] feat(SetTheory/Cardinal/Arithmetic): miscellaneous cardinality lemmas (#18933) --- Mathlib/SetTheory/Cardinal/Arithmetic.lean | 3 ++ Mathlib/SetTheory/Cardinal/Basic.lean | 35 +++++++++++++++++++++- 2 files changed, 37 insertions(+), 1 deletion(-) diff --git a/Mathlib/SetTheory/Cardinal/Arithmetic.lean b/Mathlib/SetTheory/Cardinal/Arithmetic.lean index 88177355e75c5..5a7416afca6f0 100644 --- a/Mathlib/SetTheory/Cardinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Cardinal/Arithmetic.lean @@ -260,6 +260,9 @@ theorem add_mk_eq_max {α β : Type u} [Infinite α] : #α + #β = max #α #β : theorem add_mk_eq_max' {α β : Type u} [Infinite β] : #α + #β = max #α #β := add_eq_max' (aleph0_le_mk β) +theorem add_mk_eq_self {α : Type*} [Infinite α] : #α + #α = #α := by + simp + theorem add_le_max (a b : Cardinal) : a + b ≤ max (max a b) ℵ₀ := by rcases le_or_lt ℵ₀ a with ha | ha · rw [add_eq_max ha] diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index d2919d4a8a761..e2bee0bf6130e 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -246,6 +246,11 @@ we provide this statement separately so you don't have to solve the specializati theorem lift_mk_eq' {α : Type u} {β : Type v} : lift.{v} #α = lift.{u} #β ↔ Nonempty (α ≃ β) := lift_mk_eq.{u, v, 0} +theorem mk_congr_lift {α : Type u} {β : Type v} (e : α ≃ β) : lift.{v} #α = lift.{u} #β := + lift_mk_eq'.2 ⟨e⟩ + +alias _root_.Equiv.lift_cardinal_eq := mk_congr_lift + -- Porting note: simpNF is not happy with universe levels. @[simp, nolint simpNF] theorem lift_mk_shrink (α : Type u) [Small.{v} α] : @@ -1876,6 +1881,15 @@ theorem mk_image_eq_lift {α : Type u} {β : Type v} (f : α → β) (s : Set α lift.{u} #(f '' s) = lift.{v} #s := mk_image_eq_of_injOn_lift _ _ h.injOn +@[simp] +theorem mk_image_embedding_lift {β : Type v} (f : α ↪ β) (s : Set α) : + lift.{u} #(f '' s) = lift.{v} #s := + mk_image_eq_lift _ _ f.injective + +@[simp] +theorem mk_image_embedding (f : α ↪ β) (s : Set α) : #(f '' s) = #s := by + simpa using mk_image_embedding_lift f s + theorem mk_iUnion_le_sum_mk {α ι : Type u} {f : ι → Set α} : #(⋃ i, f i) ≤ sum fun i => #(f i) := calc #(⋃ i, f i) ≤ #(Σi, f i) := mk_le_of_surjective (Set.sigmaToiUnion_surjective f) @@ -2035,6 +2049,17 @@ theorem mk_preimage_of_injective_of_subset_range (f : α → β) (s : Set β) (h (h2 : s ⊆ range f) : #(f ⁻¹' s) = #s := by convert mk_preimage_of_injective_of_subset_range_lift.{u, u} f s h h2 using 1 <;> rw [lift_id] +@[simp] +theorem mk_preimage_equiv_lift {β : Type v} (f : α ≃ β) (s : Set β) : + lift.{v} #(f ⁻¹' s) = lift.{u} #s := by + apply mk_preimage_of_injective_of_subset_range_lift _ _ f.injective + rw [f.range_eq_univ] + exact fun _ _ ↦ ⟨⟩ + +@[simp] +theorem mk_preimage_equiv (f : α ≃ β) (s : Set β) : #(f ⁻¹' s) = #s := by + simpa using mk_preimage_equiv_lift f s + theorem mk_preimage_of_injective (f : α → β) (s : Set β) (h : Injective f) : #(f ⁻¹' s) ≤ #s := by rw [← lift_id #(↑(f ⁻¹' s)), ← lift_id #(↑s)] @@ -2064,6 +2089,14 @@ theorem le_mk_iff_exists_subset {c : Cardinal} {α : Type u} {s : Set α} : rw [le_mk_iff_exists_set, ← Subtype.exists_set_subtype] apply exists_congr; intro t; rw [mk_image_eq]; apply Subtype.val_injective +@[simp] +theorem mk_range_inl {α : Type u} {β : Type v} : #(range (@Sum.inl α β)) = lift.{v} #α := by + rw [← lift_id'.{u, v} #_, (Equiv.Set.rangeInl α β).lift_cardinal_eq, lift_umax.{u, v}] + +@[simp] +theorem mk_range_inr {α : Type u} {β : Type v} : #(range (@Sum.inr α β)) = lift.{u} #β := by + rw [← lift_id'.{v, u} #_, (Equiv.Set.rangeInr α β).lift_cardinal_eq, lift_umax.{v, u}] + theorem two_le_iff : (2 : Cardinal) ≤ #α ↔ ∃ x y : α, x ≠ y := by rw [← Nat.cast_two, nat_succ, succ_le_iff, Nat.cast_one, one_lt_iff_nontrivial, nontrivial_iff] @@ -2171,4 +2204,4 @@ end Cardinal -- end Tactic -set_option linter.style.longFile 2200 +set_option linter.style.longFile 2400 From 53f6c026d38b779af07930c4ea6467c6ab3c0dd2 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Sun, 5 Jan 2025 07:56:23 +0000 Subject: [PATCH 152/189] chore: update Mathlib dependencies 2025-01-05 (#20485) 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 8a7f0e3ba49a2..795e5c8c6acca 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "8ce422eb59adf557fac184f8b1678c75fa03075c", + "rev": "b2e8b6868397fcd93c00aca7278b933c16c0ffb3", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "v4.16.0-rc1", From b552a67b9245b0012c186a4a06f519600b151e28 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Sun, 5 Jan 2025 08:51:50 +0000 Subject: [PATCH 153/189] fix: Stop cancelling builds of master (#20435) As mentioned on [zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Persistent.20.23allow_unused_tactic/near/491740884), the building process on `master` can get cancelled on occasion. this PR hopes to fix that issue. --- .github/build.in.yml | 19 +++++++------------ .github/workflows/bors.yml | 19 +++++++------------ .github/workflows/build.yml | 19 +++++++------------ .github/workflows/build_fork.yml | 19 +++++++------------ 4 files changed, 28 insertions(+), 48 deletions(-) diff --git a/.github/build.in.yml b/.github/build.in.yml index c56faf4061fea..b7acc84572050 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -12,19 +12,14 @@ env: # not necessarily satisfy this property. LAKE_NO_CACHE: true -jobs: - # Cancels previous runs of jobs in this file - cancel: - if: github.repository == 'leanprover-community/mathlib4' - name: 'Cancel Previous Runs (CI)' - runs-on: ubuntu-latest - # timeout-minutes: 3 - steps: - - uses: styfle/cancel-workflow-action@0.12.1 - with: - all_but_latest: true - access_token: ${{ github.token }} +concurrency: + # label each workflow run; only the latest with each label will run + # workflows on master get more expressive labels + group: ${{ github.workflow }}-${{ github.ref }}.${{(github.ref == 'refs/heads/master' && github.run_id) || ''}} + # cancel any running workflow with the same label + cancel-in-progress: true +jobs: style_lint: if: github.repository MAIN_OR_FORK 'leanprover-community/mathlib4' name: Lint styleJOB_NAME diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 869b4d379c959..37c0ac50d19a0 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -22,19 +22,14 @@ env: # not necessarily satisfy this property. LAKE_NO_CACHE: true -jobs: - # Cancels previous runs of jobs in this file - cancel: - if: github.repository == 'leanprover-community/mathlib4' - name: 'Cancel Previous Runs (CI)' - runs-on: ubuntu-latest - # timeout-minutes: 3 - steps: - - uses: styfle/cancel-workflow-action@0.12.1 - with: - all_but_latest: true - access_token: ${{ github.token }} +concurrency: + # label each workflow run; only the latest with each label will run + # workflows on master get more expressive labels + group: ${{ github.workflow }}-${{ github.ref }}.${{(github.ref == 'refs/heads/master' && github.run_id) || ''}} + # cancel any running workflow with the same label + cancel-in-progress: true +jobs: style_lint: if: github.repository == 'leanprover-community/mathlib4' name: Lint style diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 25c242698e644..62525af3c3089 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -29,19 +29,14 @@ env: # not necessarily satisfy this property. LAKE_NO_CACHE: true -jobs: - # Cancels previous runs of jobs in this file - cancel: - if: github.repository == 'leanprover-community/mathlib4' - name: 'Cancel Previous Runs (CI)' - runs-on: ubuntu-latest - # timeout-minutes: 3 - steps: - - uses: styfle/cancel-workflow-action@0.12.1 - with: - all_but_latest: true - access_token: ${{ github.token }} +concurrency: + # label each workflow run; only the latest with each label will run + # workflows on master get more expressive labels + group: ${{ github.workflow }}-${{ github.ref }}.${{(github.ref == 'refs/heads/master' && github.run_id) || ''}} + # cancel any running workflow with the same label + cancel-in-progress: true +jobs: style_lint: if: github.repository == 'leanprover-community/mathlib4' name: Lint style diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 9848f1817b72d..14ae2662c9d09 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -26,19 +26,14 @@ env: # not necessarily satisfy this property. LAKE_NO_CACHE: true -jobs: - # Cancels previous runs of jobs in this file - cancel: - if: github.repository == 'leanprover-community/mathlib4' - name: 'Cancel Previous Runs (CI)' - runs-on: ubuntu-latest - # timeout-minutes: 3 - steps: - - uses: styfle/cancel-workflow-action@0.12.1 - with: - all_but_latest: true - access_token: ${{ github.token }} +concurrency: + # label each workflow run; only the latest with each label will run + # workflows on master get more expressive labels + group: ${{ github.workflow }}-${{ github.ref }}.${{(github.ref == 'refs/heads/master' && github.run_id) || ''}} + # cancel any running workflow with the same label + cancel-in-progress: true +jobs: style_lint: if: github.repository != 'leanprover-community/mathlib4' name: Lint style (fork) From a1c721608e3004545cfb278df2c13a87d1862345 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sun, 5 Jan 2025 11:06:22 +0000 Subject: [PATCH 154/189] =?UTF-8?q?feat:=20add=20`IsStarNormal=20(?= =?UTF-8?q?=E2=86=91x=E2=81=BB=C2=B9=20:=20R)`=20instance=20for=20`x=20:?= =?UTF-8?q?=20R=CB=A3`=20(#20091)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Also renames `isStarNormal_{zero, one}` to `IsStarNormal.{zero, one}`. --- Mathlib/Algebra/Star/SelfAdjoint.lean | 8 ++++++-- .../ContinuousFunctionalCalculus/Instances.lean | 2 +- 2 files changed, 7 insertions(+), 3 deletions(-) diff --git a/Mathlib/Algebra/Star/SelfAdjoint.lean b/Mathlib/Algebra/Star/SelfAdjoint.lean index 2acc10366d170..b43afd4b955f9 100644 --- a/Mathlib/Algebra/Star/SelfAdjoint.lean +++ b/Mathlib/Algebra/Star/SelfAdjoint.lean @@ -563,10 +563,10 @@ theorem isSelfAdjoint_smul_of_mem_skewAdjoint [Ring R] [AddCommGroup A] [Module (ha : a ∈ skewAdjoint A) : IsSelfAdjoint (r • a) := (star_smul _ _).trans <| (congr_arg₂ _ hr ha).trans <| neg_smul_neg _ _ -instance isStarNormal_zero [Semiring R] [StarRing R] : IsStarNormal (0 : R) := +protected instance IsStarNormal.zero [Semiring R] [StarRing R] : IsStarNormal (0 : R) := ⟨by simp only [Commute.refl, star_comm_self, star_zero]⟩ -instance isStarNormal_one [MulOneClass R] [StarMul R] : IsStarNormal (1 : R) := +protected instance IsStarNormal.one [MulOneClass R] [StarMul R] : IsStarNormal (1 : R) := ⟨by simp only [Commute.refl, star_comm_self, star_one]⟩ protected instance IsStarNormal.star [Mul R] [StarMul R] {x : R} [IsStarNormal x] : @@ -577,6 +577,10 @@ protected instance IsStarNormal.neg [Ring R] [StarAddMonoid R] {x : R} [IsStarNo IsStarNormal (-x) := ⟨show star (-x) * -x = -x * star (-x) by simp_rw [star_neg, neg_mul_neg, star_comm_self']⟩ +protected instance IsStarNormal.val_inv [Monoid R] [StarMul R] {x : Rˣ} [IsStarNormal (x : R)] : + IsStarNormal (↑x⁻¹ : R) where + star_comm_self := by simpa [← Units.coe_star_inv, -Commute.units_val_iff] using star_comm_self + protected instance IsStarNormal.map {F R S : Type*} [Mul R] [Star R] [Mul S] [Star S] [FunLike F R S] [MulHomClass F R S] [StarHomClass F R S] (f : F) (r : R) [hr : IsStarNormal r] : IsStarNormal (f r) where diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index 8e43434bef5ac..3fe783ba69ee2 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -155,7 +155,7 @@ section Normal instance IsStarNormal.instContinuousFunctionalCalculus {A : Type*} [CStarAlgebra A] : ContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop) where - predicate_zero := isStarNormal_zero + predicate_zero := .zero spectrum_nonempty a _ := spectrum.nonempty a exists_cfc_of_predicate a ha := by refine ⟨(StarAlgebra.elemental ℂ a).subtype.comp <| continuousFunctionalCalculus a, From f1f1e33b53f42745b770476050ace133ea03d764 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sun, 5 Jan 2025 11:29:37 +0000 Subject: [PATCH 155/189] chore(CategoryTheory): make relevant parameters explicit in `Arrow.homMk`. (#20259) The two morphisms `left` and `right` which constitute the data in a morphism in the category `Arrow C` of arrows are made explicit, and the factorization property gets a default value `by aesop_cat`. --- Mathlib/Algebra/Homology/HomologicalComplex.lean | 4 ++-- Mathlib/CategoryTheory/Comma/Arrow.lean | 7 ++++--- Mathlib/CategoryTheory/Limits/Shapes/Images.lean | 4 ++-- Mathlib/CategoryTheory/Square.lean | 12 ++++++------ 4 files changed, 14 insertions(+), 13 deletions(-) diff --git a/Mathlib/Algebra/Homology/HomologicalComplex.lean b/Mathlib/Algebra/Homology/HomologicalComplex.lean index 03b17997ea652..9fc057de9025d 100644 --- a/Mathlib/Algebra/Homology/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/HomologicalComplex.lean @@ -566,7 +566,7 @@ attribute [simp] comm_to_apply induces a morphism of arrows of the differentials out of each object. -/ def sqFrom (f : Hom C₁ C₂) (i : ι) : Arrow.mk (C₁.dFrom i) ⟶ Arrow.mk (C₂.dFrom i) := - Arrow.homMk (f.comm_from i) + Arrow.homMk _ _ (f.comm_from i) @[simp] theorem sqFrom_left (f : Hom C₁ C₂) (i : ι) : (f.sqFrom i).left = f.f i := @@ -589,7 +589,7 @@ theorem sqFrom_comp (f : C₁ ⟶ C₂) (g : C₂ ⟶ C₃) (i : ι) : induces a morphism of arrows of the differentials into each object. -/ def sqTo (f : Hom C₁ C₂) (j : ι) : Arrow.mk (C₁.dTo j) ⟶ Arrow.mk (C₂.dTo j) := - Arrow.homMk (f.comm_to j) + Arrow.homMk _ _ (f.comm_to j) @[simp] theorem sqTo_left (f : Hom C₁ C₂) (j : ι) : (f.sqTo j).left = f.prev j := diff --git a/Mathlib/CategoryTheory/Comma/Arrow.lean b/Mathlib/CategoryTheory/Comma/Arrow.lean index cc6d53e196a94..6a161770de191 100644 --- a/Mathlib/CategoryTheory/Comma/Arrow.lean +++ b/Mathlib/CategoryTheory/Comma/Arrow.lean @@ -94,15 +94,16 @@ instance {X Y : T} : CoeOut (X ⟶ Y) (Arrow T) where /-- A morphism in the arrow category is a commutative square connecting two objects of the arrow category. -/ @[simps] -def homMk {f g : Arrow T} {u : f.left ⟶ g.left} {v : f.right ⟶ g.right} - (w : u ≫ g.hom = f.hom ≫ v) : f ⟶ g where +def homMk {f g : Arrow T} (u : f.left ⟶ g.left) (v : f.right ⟶ g.right) + (w : u ≫ g.hom = f.hom ≫ v := by aesop_cat) : f ⟶ g where left := u right := v w := w /-- We can also build a morphism in the arrow category out of any commutative square in `T`. -/ @[simps] -def homMk' {X Y : T} {f : X ⟶ Y} {P Q : T} {g : P ⟶ Q} {u : X ⟶ P} {v : Y ⟶ Q} (w : u ≫ g = f ≫ v) : +def homMk' {X Y : T} {f : X ⟶ Y} {P Q : T} {g : P ⟶ Q} (u : X ⟶ P) (v : Y ⟶ Q) + (w : u ≫ g = f ≫ v := by aesop_cat) : Arrow.mk f ⟶ Arrow.mk g where left := u right := v diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean index fec3f30cad1ff..1214bfe40be59 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean @@ -717,8 +717,8 @@ theorem image.factor_map : theorem image.map_ι : image.map sq ≫ image.ι g.hom = image.ι f.hom ≫ sq.right := by simp theorem image.map_homMk'_ι {X Y P Q : C} {k : X ⟶ Y} [HasImage k] {l : P ⟶ Q} [HasImage l] - {m : X ⟶ P} {n : Y ⟶ Q} (w : m ≫ l = k ≫ n) [HasImageMap (Arrow.homMk' w)] : - image.map (Arrow.homMk' w) ≫ image.ι l = image.ι k ≫ n := + {m : X ⟶ P} {n : Y ⟶ Q} (w : m ≫ l = k ≫ n) [HasImageMap (Arrow.homMk' _ _ w)] : + image.map (Arrow.homMk' _ _ w) ≫ image.ι l = image.ι k ≫ n := image.map_ι _ section diff --git a/Mathlib/CategoryTheory/Square.lean b/Mathlib/CategoryTheory/Square.lean index 1286c1bd87b9b..9062feb1d215c 100644 --- a/Mathlib/CategoryTheory/Square.lean +++ b/Mathlib/CategoryTheory/Square.lean @@ -172,9 +172,9 @@ commutative square `sq` to the obvious arrow from the left morphism of `sq` to the right morphism of `sq`. -/ @[simps!] def toArrowArrowFunctor : Square C ⥤ Arrow (Arrow C) where - obj sq := Arrow.mk (Arrow.homMk sq.fac : Arrow.mk sq.f₁₃ ⟶ Arrow.mk sq.f₂₄) - map φ := Arrow.homMk (u := Arrow.homMk φ.comm₁₃.symm) - (v := Arrow.homMk φ.comm₂₄.symm) (by aesop_cat) + obj sq := Arrow.mk (Arrow.homMk _ _ sq.fac : Arrow.mk sq.f₁₃ ⟶ Arrow.mk sq.f₂₄) + map φ := Arrow.homMk (Arrow.homMk _ _ φ.comm₁₃.symm) + (Arrow.homMk _ _ φ.comm₂₄.symm) /-- The functor `Arrow (Arrow C) ⥤ Square C` which sends a morphism `Arrow.mk f ⟶ Arrow.mk g` to the commutative square @@ -207,9 +207,9 @@ commutative square `sq` to the obvious arrow from the top morphism of `sq` to the bottom morphism of `sq`. -/ @[simps!] def toArrowArrowFunctor' : Square C ⥤ Arrow (Arrow C) where - obj sq := Arrow.mk (Arrow.homMk sq.fac.symm : Arrow.mk sq.f₁₂ ⟶ Arrow.mk sq.f₃₄) - map φ := Arrow.homMk (u := Arrow.homMk φ.comm₁₂.symm) - (v := Arrow.homMk φ.comm₃₄.symm) (by aesop_cat) + obj sq := Arrow.mk (Arrow.homMk _ _ sq.fac.symm : Arrow.mk sq.f₁₂ ⟶ Arrow.mk sq.f₃₄) + map φ := Arrow.homMk (Arrow.homMk _ _ φ.comm₁₂.symm) + (Arrow.homMk _ _ φ.comm₃₄.symm) /-- The functor `Arrow (Arrow C) ⥤ Square C` which sends a morphism `Arrow.mk f ⟶ Arrow.mk g` to the commutative square From 8a6db46485c691c5bb94154d3dc218f99277c723 Mon Sep 17 00:00:00 2001 From: "Tristan F.-R." Date: Sun, 5 Jan 2025 11:29:38 +0000 Subject: [PATCH 156/189] chore(Combinatorics/Enumerative/Partition): auto-derive DecidableEq (#20277) Original bug fixed in https://github.com/leanprover/lean4/pull/2918. Co-authored-by: Tristan F. --- Mathlib/Combinatorics/Enumerative/Partition.lean | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index b2226025ed657..bdc8edeb9c657 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -54,12 +54,11 @@ structure Partition (n : ℕ) where parts_pos : ∀ {i}, i ∈ parts → 0 < i /-- proof that the `parts` sum to `n`-/ parts_sum : parts.sum = n - -- Porting note: chokes on `parts_pos` - --deriving DecidableEq +deriving DecidableEq namespace Partition --- TODO: This should be automatically derived, see https://github.com/leanprover/lean4/issues/2914 +@[deprecated "Partition now derives an instance of DecidableEq." (since := "2024-12-28")] instance decidableEqPartition {n : ℕ} : DecidableEq (Partition n) := fun _ _ => decidable_of_iff' _ Partition.ext_iff From f916a16fa7a15fea30540ebafac0e7e8dbf3b97b Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Sun, 5 Jan 2025 11:29:40 +0000 Subject: [PATCH 157/189] chore(MvPolynomial/Localization): golf using TensorProduct (#20309) --- .../RingTheory/MvPolynomial/Localization.lean | 41 +++---------------- 1 file changed, 6 insertions(+), 35 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Localization.lean b/Mathlib/RingTheory/MvPolynomial/Localization.lean index 2a8977ca29026..e3fb28794c4d7 100644 --- a/Mathlib/RingTheory/MvPolynomial/Localization.lean +++ b/Mathlib/RingTheory/MvPolynomial/Localization.lean @@ -3,11 +3,13 @@ Copyright (c) 2024 Christian Merten. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Christian Merten -/ +import Mathlib.Algebra.Module.LocalizedModule.IsLocalization import Mathlib.Algebra.MvPolynomial.CommRing import Mathlib.RingTheory.Ideal.Quotient.Operations import Mathlib.RingTheory.Localization.Away.Basic -import Mathlib.RingTheory.Localization.Basic +import Mathlib.RingTheory.Localization.BaseChange import Mathlib.RingTheory.MvPolynomial.Basic +import Mathlib.RingTheory.TensorProduct.MvPolynomial /-! @@ -36,40 +38,9 @@ attribute [local instance] algebraMvPolynomial If `S` is the localization of `R` at a submonoid `M`, then `MvPolynomial σ S` is the localization of `MvPolynomial σ R` at `M.map MvPolynomial.C`. -/ -instance isLocalization : IsLocalization (M.map <| C (σ := σ)) - (MvPolynomial σ S) where - map_units' := by - rintro ⟨p, q, hq, rfl⟩ - simp only [algebraMap_def, map_C] - exact IsUnit.map _ (IsLocalization.map_units _ ⟨q, hq⟩) - surj' p := by - simp only [algebraMap_def, Prod.exists, Subtype.exists, - Submonoid.mem_map, exists_prop, exists_exists_and_eq_and, map_C] - refine induction_on' p ?_ ?_ - · intro u s - obtain ⟨⟨r, m⟩, hr⟩ := IsLocalization.surj M s - use monomial u r, m, m.property - simp only [map_monomial] - rw [← hr, mul_comm, C_mul_monomial, mul_comm] - · intro p p' ⟨x, m, hm, hxm⟩ ⟨x', m', hm', hxm'⟩ - use x * (C m') + x' * (C m), m * m', Submonoid.mul_mem _ hm hm' - simp only [map_mul, map_add, map_C] - rw [add_mul, ← mul_assoc, hxm, ← mul_assoc, ← hxm, ← hxm'] - ring - exists_of_eq {p q} := by - intro h - 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 - simp only [coeff_C_mul] - by_cases h : m ∈ p.support ∪ q.support - · exact Finset.prod_mul_eq_prod_mul_of_exists m h (hc m) - · simp only [Finset.mem_union, mem_support_iff, ne_eq, not_or, Decidable.not_not] at h - rw [h.left, h.right] +instance isLocalization : IsLocalization (M.map <| C (σ := σ)) (MvPolynomial σ S) := + isLocalizedModule_iff_isLocalization.mp <| (isLocalizedModule_iff_isBaseChange M S _).mpr <| + .of_equiv (algebraTensorAlgEquiv _ _).toLinearEquiv fun _ ↦ by simp lemma isLocalization_C_mk' (a : R) (m : M) : C (IsLocalization.mk' S a m) = IsLocalization.mk' (MvPolynomial σ S) (C (σ := σ) a) From 238e9b996a074042502c34865b46ebbf0adcf44f Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 5 Jan 2025 11:29:41 +0000 Subject: [PATCH 158/189] chore: generalize `FormalMultilinearSeries.ofScalars_norm` to `Seminormed` (#20351) More usefully, this does the same for a collection of results about `mkPiAlgebraFin`. --- Mathlib/Analysis/Analytic/OfScalars.lean | 17 ++++++++++++++--- .../Analysis/NormedSpace/Multilinear/Basic.lean | 2 +- 2 files changed, 15 insertions(+), 4 deletions(-) diff --git a/Mathlib/Analysis/Analytic/OfScalars.lean b/Mathlib/Analysis/Analytic/OfScalars.lean index 771e16242e8ad..5674b07141092 100644 --- a/Mathlib/Analysis/Analytic/OfScalars.lean +++ b/Mathlib/Analysis/Analytic/OfScalars.lean @@ -139,12 +139,12 @@ theorem ofScalarsSum_unop [T2Space E] (x : Eᵐᵒᵖ) : end Field -section Normed +section Seminormed open Filter ENNReal open scoped Topology NNReal -variable {𝕜 : Type*} (E : Type*) [NontriviallyNormedField 𝕜] [NormedRing E] +variable {𝕜 : Type*} (E : Type*) [NontriviallyNormedField 𝕜] [SeminormedRing E] [NormedAlgebra 𝕜 E] (c : ℕ → 𝕜) (n : ℕ) -- Also works: @@ -152,7 +152,8 @@ variable {𝕜 : Type*} (E : Type*) [NontriviallyNormedField 𝕜] [NormedRing E set_option maxSynthPendingDepth 2 in theorem ofScalars_norm_eq_mul : ‖ofScalars E c n‖ = ‖c n‖ * ‖ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 n E‖ := by - rw [ofScalars, norm_smul (c n) (ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 n E)] + set_option maxSynthPendingDepth 2 in + rw [ofScalars, norm_smul] theorem ofScalars_norm_le (hn : n > 0) : ‖ofScalars E c n‖ ≤ ‖c n‖ := by simp only [ofScalars_norm_eq_mul] @@ -163,6 +164,16 @@ theorem ofScalars_norm_le (hn : n > 0) : ‖ofScalars E c n‖ ≤ ‖c n‖ := theorem ofScalars_norm [NormOneClass E] : ‖ofScalars E c n‖ = ‖c n‖ := by simp [ofScalars_norm_eq_mul] +end Seminormed + +section Normed + +open Filter ENNReal +open scoped Topology NNReal + +variable {𝕜 : Type*} (E : Type*) [NontriviallyNormedField 𝕜] [NormedRing E] + [NormedAlgebra 𝕜 E] (c : ℕ → 𝕜) (n : ℕ) + private theorem tendsto_succ_norm_div_norm {r r' : ℝ≥0} (hr' : r' ≠ 0) (hc : Tendsto (fun n ↦ ‖c n.succ‖ / ‖c n‖) atTop (𝓝 r)) : Tendsto (fun n ↦ ‖‖c (n + 1)‖ * r' ^ (n + 1)‖ / diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 6c60138a2473d..4a71d1dff9424 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -785,7 +785,7 @@ end section -variable {n : ℕ} {A : Type*} [NormedRing A] [NormedAlgebra 𝕜 A] +variable {n : ℕ} {A : Type*} [SeminormedRing A] [NormedAlgebra 𝕜 A] theorem norm_mkPiAlgebraFin_succ_le : ‖ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 n.succ A‖ ≤ 1 := by refine opNorm_le_bound zero_le_one fun m => ?_ From f82d0858596934d61137b67b5e209beeca0db408 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 5 Jan 2025 11:29:42 +0000 Subject: [PATCH 159/189] =?UTF-8?q?chore(SetTheory/Ordinal/Basic):=20`{x?= =?UTF-8?q?=20//=20x=20<=20y}`=20=E2=86=92=20`Iio=20y`=20(#20413)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 8 +------- Mathlib/SetTheory/Ordinal/Basic.lean | 12 +++++------- 2 files changed, 6 insertions(+), 14 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index ee879a0ea1dde..c914bbd405ef9 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -338,7 +338,7 @@ theorem enum_succ_eq_top {o : Ordinal} : theorem has_succ_of_type_succ_lt {α} {r : α → α → Prop} [wo : IsWellOrder α r] (h : ∀ a < type r, succ a < type r) (x : α) : ∃ y, r x y := by use enum r ⟨succ (typein r x), h _ (typein_lt_type r x)⟩ - convert enum_lt_enum (o₁ := ⟨_, typein_lt_type r x⟩) (o₂ := ⟨_, h _ (typein_lt_type r x)⟩).mpr _ + convert enum_lt_enum.mpr _ · rw [enum_typein] · rw [Subtype.mk_lt_mk, lt_succ_iff] @@ -364,12 +364,6 @@ theorem typein_ordinal (o : Ordinal.{u}) : rintro ⟨α, r, wo⟩; apply Quotient.sound constructor; refine ((RelIso.preimage Equiv.ulift r).trans (enum r).symm).symm --- Porting note: `· < ·` requires a type ascription for an `IsWellOrder` instance. -@[deprecated typein_ordinal (since := "2024-09-19")] -theorem type_subrel_lt (o : Ordinal.{u}) : - type (@Subrel Ordinal (· < ·) { o' : Ordinal | o' < o }) = Ordinal.lift.{u + 1} o := - typein_ordinal o - theorem mk_Iio_ordinal (o : Ordinal.{u}) : #(Iio o) = Cardinal.lift.{u + 1} o.card := by rw [lift_card, ← typein_ordinal] diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index a663b8dadc323..33a1d3e36a834 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -487,10 +487,8 @@ theorem typein_surjOn (r : α → α → Prop) [IsWellOrder α r] : That is, `enum` maps an initial segment of the ordinals, those less than the order type of `r`, to the elements of `α`. -/ --- The explicit typing is required in order for `simp` to work properly. @[simps! symm_apply_coe] -def enum (r : α → α → Prop) [IsWellOrder α r] : - @RelIso { o // o < type r } α (Subrel (· < ·) { o | o < type r }) r := +def enum (r : α → α → Prop) [IsWellOrder α r] : (· < · : Iio (type r) → Iio (type r) → Prop) ≃r r := (typein r).subrelIso @[simp] @@ -507,20 +505,20 @@ theorem enum_typein (r : α → α → Prop) [IsWellOrder α r] (a : α) : enum r ⟨typein r a, typein_lt_type r a⟩ = a := enum_type (PrincipalSeg.ofElement r a) -theorem enum_lt_enum {r : α → α → Prop} [IsWellOrder α r] {o₁ o₂ : {o // o < type r}} : +theorem enum_lt_enum {r : α → α → Prop} [IsWellOrder α r] {o₁ o₂ : Iio (type r)} : r (enum r o₁) (enum r o₂) ↔ o₁ < o₂ := (enum _).map_rel_iff -theorem enum_le_enum (r : α → α → Prop) [IsWellOrder α r] {o₁ o₂ : {o // o < type r}} : +theorem enum_le_enum (r : α → α → Prop) [IsWellOrder α r] {o₁ o₂ : Iio (type r)} : ¬r (enum r o₁) (enum r o₂) ↔ o₂ ≤ o₁ := by rw [enum_lt_enum (r := r), not_lt] @[simp] -theorem enum_le_enum' (a : Ordinal) {o₁ o₂ : {o // o < type (· < ·)}} : +theorem enum_le_enum' (a : Ordinal) {o₁ o₂ : Iio (type (· < ·))} : enum (· < ·) o₁ ≤ enum (α := a.toType) (· < ·) o₂ ↔ o₁ ≤ o₂ := by rw [← enum_le_enum, not_lt] -theorem enum_inj {r : α → α → Prop} [IsWellOrder α r] {o₁ o₂ : {o // o < type r}} : +theorem enum_inj {r : α → α → Prop} [IsWellOrder α r] {o₁ o₂ : Iio (type r)} : enum r o₁ = enum r o₂ ↔ o₁ = o₂ := EmbeddingLike.apply_eq_iff_eq _ From aef9ea68f33362fb3c434fc1064829b2ccd9653b Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 5 Jan 2025 11:29:43 +0000 Subject: [PATCH 160/189] feat(FaaDiBruno): derive `DecidableEq` (#20482) I'm working on a formula for the Taylor series of the inverse function for the Sard-Moreira project, and I need complement to work for `Finset (OrderedFinpartition _)`. --- Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean index 0ebab41ba1c2a..eb2240b7ca333 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean @@ -106,6 +106,7 @@ structure OrderedFinpartition (n : ℕ) where disjoint : PairwiseDisjoint univ fun m ↦ range (emb m) /-- The parts cover everything -/ cover x : ∃ m, x ∈ range (emb m) + deriving DecidableEq namespace OrderedFinpartition From 316642556f64cc298ff78aba13c5e681833d1c10 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Sun, 5 Jan 2025 11:59:22 +0000 Subject: [PATCH 161/189] chore: drop redundant hypothesis in `Module.Dual.eq_of_preReflection_mapsTo` (#20492) --- Mathlib/LinearAlgebra/Reflection.lean | 10 +++++----- Mathlib/LinearAlgebra/RootSystem/Basic.lean | 8 ++++---- 2 files changed, 9 insertions(+), 9 deletions(-) diff --git a/Mathlib/LinearAlgebra/Reflection.lean b/Mathlib/LinearAlgebra/Reflection.lean index 7201b2603f574..612afa9e70760 100644 --- a/Mathlib/LinearAlgebra/Reflection.lean +++ b/Mathlib/LinearAlgebra/Reflection.lean @@ -311,10 +311,11 @@ This rather technical-looking lemma exists because it is exactly what is needed uniqueness results for root data / systems. One might regard this lemma as lying at the boundary of linear algebra and combinatorics since the finiteness assumption is the key. -/ lemma Dual.eq_of_preReflection_mapsTo [CharZero R] [NoZeroSMulDivisors R M] - {x : M} (hx : x ≠ 0) {Φ : Set M} (hΦ₁ : Φ.Finite) (hΦ₂ : span R Φ = ⊤) {f g : Dual R M} + {x : M} {Φ : Set M} (hΦ₁ : Φ.Finite) (hΦ₂ : span R Φ = ⊤) {f g : Dual R M} (hf₁ : f x = 2) (hf₂ : MapsTo (preReflection x f) Φ Φ) (hg₁ : g x = 2) (hg₂ : MapsTo (preReflection x g) Φ Φ) : f = g := by + have hx : x ≠ 0 := by rintro rfl; simp at hf₁ let u := reflection hg₁ * reflection hf₁ have hu : u = LinearMap.id (R := R) (M := M) + (f - g).smulRight x := by ext y @@ -344,7 +345,7 @@ lemma Dual.eq_of_preReflection_mapsTo [CharZero R] [NoZeroSMulDivisors R M] uniqueness result for root data. See the doc string of `Module.Dual.eq_of_preReflection_mapsTo` for further remarks. -/ lemma Dual.eq_of_preReflection_mapsTo' [CharZero R] [NoZeroSMulDivisors R M] - {x : M} (hx : x ≠ 0) {Φ : Set M} (hΦ₁ : Φ.Finite) (hx' : x ∈ span R Φ) {f g : Dual R M} + {x : M} {Φ : Set M} (hΦ₁ : Φ.Finite) (hx : x ∈ span R Φ) {f g : Dual R M} (hf₁ : f x = 2) (hf₂ : MapsTo (preReflection x f) Φ Φ) (hg₁ : g x = 2) (hg₂ : MapsTo (preReflection x g) Φ Φ) : (span R Φ).subtype.dualMap f = (span R Φ).subtype.dualMap g := by @@ -355,8 +356,7 @@ lemma Dual.eq_of_preReflection_mapsTo' [CharZero R] [NoZeroSMulDivisors R M] simp only [Φ'] rw [range_inclusion] simp - let x' : span R Φ := ⟨x, hx'⟩ - have hx' : x' ≠ 0 := Subtype.coe_ne_coe.1 hx + let x' : span R Φ := ⟨x, hx⟩ have this : ∀ {F : Dual R M}, MapsTo (preReflection x F) Φ Φ → MapsTo (preReflection x' ((span R Φ).subtype.dualMap F)) Φ' Φ' := by intro F hF ⟨y, hy⟩ hy' @@ -365,7 +365,7 @@ lemma Dual.eq_of_preReflection_mapsTo' [CharZero R] [NoZeroSMulDivisors R M] simp only [SetLike.coe_sort_coe, mem_setOf_eq] at hy' ⊢ rw [range_inclusion] exact hF hy' - exact eq_of_preReflection_mapsTo hx' hΦ'₁ hΦ'₂ hf₁ (this hf₂) hg₁ (this hg₂) + exact eq_of_preReflection_mapsTo hΦ'₁ hΦ'₂ hf₁ (this hf₂) hg₁ (this hg₂) variable {y} variable {g : Dual R M} diff --git a/Mathlib/LinearAlgebra/RootSystem/Basic.lean b/Mathlib/LinearAlgebra/RootSystem/Basic.lean index ba9290e8307a8..749ec3364bcb0 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Basic.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Basic.lean @@ -135,7 +135,7 @@ protected lemma ext [CharZero R] [NoZeroSMulDivisors R M] ext i apply P₁.injOn_dualMap_subtype_span_root_coroot (mem_range_self i) (hc ▸ mem_range_self i) simp only [LinearMap.coe_comp, LinearEquiv.coe_coe, comp_apply] - apply Dual.eq_of_preReflection_mapsTo' (P₁.ne_zero i) (finite_range P₁.root) + apply Dual.eq_of_preReflection_mapsTo' (finite_range P₁.root) · exact Submodule.subset_span (mem_range_self i) · exact P₁.coroot_root_two i · exact P₁.mapsTo_reflection_root i @@ -177,7 +177,7 @@ private lemma coroot_eq_coreflection_of_root_eq' [CharZero R] [NoZeroSMulDivisor have := injOn_dualMap_subtype_span_range_range (finite_range root) (c := p.flip ∘ coroot) hp hr apply this (mem_range_self k) (mem_range_self l) - refine Dual.eq_of_preReflection_mapsTo' hk₀ (finite_range root) + refine Dual.eq_of_preReflection_mapsTo' (finite_range root) (Submodule.subset_span <| mem_range_self k) (hp k) (hr k) hkl ?_ rw [comp_apply, hl, hk, hij] exact (hr i).comp <| (hr j).comp (hr i) @@ -232,7 +232,7 @@ protected lemma ext [CharZero R] [NoZeroSMulDivisors R M] rintro P₁ P₂ he hr - ⟨i, rfl⟩ use i apply P₁.bijectiveRight.injective - apply Dual.eq_of_preReflection_mapsTo (P₁.ne_zero i) (finite_range P₁.root) P₁.span_eq_top + apply Dual.eq_of_preReflection_mapsTo (finite_range P₁.root) P₁.span_eq_top · exact hr ▸ he ▸ P₂.coroot_root_two i · exact hr ▸ he ▸ P₂.mapsTo_reflection_root i · exact P₁.coroot_root_two i @@ -261,7 +261,7 @@ private lemma coroot_eq_coreflection_of_root_eq_of_span_eq_top [CharZero R] [NoZ preReflection_apply] -- v4.7.0-rc1 issues have hk₀ : root k ≠ 0 := fun h ↦ by simpa [h, ← PerfectPairing.toLin_apply] using hp k apply p.bijectiveRight.injective - apply Dual.eq_of_preReflection_mapsTo hk₀ (finite_range root) hsp (hp k) (hs k) + apply Dual.eq_of_preReflection_mapsTo (finite_range root) hsp (hp k) (hs k) · simp [map_sub, α, β, α', β', sα, sβ, sα', hk, preReflection_apply, hp i, hp j, mul_two, mul_comm (p α β')] ring -- v4.7.0-rc1 issues From 1d1c8c25c580fac9b5121656d844a8d10f1c08b7 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 5 Jan 2025 12:29:59 +0000 Subject: [PATCH 162/189] feat: `Ico`, `Ioc`, and `Ioo` are not closed or compact (#20479) Inspired by [this Zulip thread](https://leanprover.zulipchat.com/#narrow/channel/113489-new-members/topic/Proving.20compactness.20of.20.280.2Cr.5D.20from.20.5B0.2Cr.5D/near/491839141) which tried to prove a false result. Also includes 12 stupid lemmas about interval equalities to make `simp` slightly more helpful. I don't know if the topological results can be generalized further, but these lemmas at least work for `Real` as was needed on Zulip. --- Mathlib/Order/Interval/Set/Basic.lean | 49 ++++++++++++++++++++++ Mathlib/Topology/Order/Compact.lean | 20 +++++++++ Mathlib/Topology/Order/DenselyOrdered.lean | 27 +++++++++++- 3 files changed, 94 insertions(+), 2 deletions(-) diff --git a/Mathlib/Order/Interval/Set/Basic.lean b/Mathlib/Order/Interval/Set/Basic.lean index fb3aefe5d8181..06d3ef45f57c2 100644 --- a/Mathlib/Order/Interval/Set/Basic.lean +++ b/Mathlib/Order/Interval/Set/Basic.lean @@ -492,6 +492,54 @@ theorem not_mem_Ioo_of_le (ha : c ≤ a) : c ∉ Ioo a b := fun h => lt_irrefl _ theorem not_mem_Ioo_of_ge (hb : b ≤ c) : c ∉ Ioo a b := fun h => lt_irrefl _ <| h.2.trans_le hb +section matched_intervals + +@[simp] theorem Icc_eq_Ioc_same_iff : Icc a b = Ioc a b ↔ ¬a ≤ b where + mp h := by simpa using Set.ext_iff.mp h a + mpr h := by rw [Icc_eq_empty h, Ioc_eq_empty (mt le_of_lt h)] + +@[simp] theorem Icc_eq_Ico_same_iff : Icc a b = Ico a b ↔ ¬a ≤ b where + mp h := by simpa using Set.ext_iff.mp h b + mpr h := by rw [Icc_eq_empty h, Ico_eq_empty (mt le_of_lt h)] + +@[simp] theorem Icc_eq_Ioo_same_iff : Icc a b = Ioo a b ↔ ¬a ≤ b where + mp h := by simpa using Set.ext_iff.mp h b + mpr h := by rw [Icc_eq_empty h, Ioo_eq_empty (mt le_of_lt h)] + +@[simp] theorem Ioc_eq_Ico_same_iff : Ioc a b = Ico a b ↔ ¬a < b where + mp h := by simpa using Set.ext_iff.mp h a + mpr h := by rw [Ioc_eq_empty h, Ico_eq_empty h] + +@[simp] theorem Ioo_eq_Ioc_same_iff : Ioo a b = Ioc a b ↔ ¬a < b where + mp h := by simpa using Set.ext_iff.mp h b + mpr h := by rw [Ioo_eq_empty h, Ioc_eq_empty h] + +@[simp] theorem Ioo_eq_Ico_same_iff : Ioo a b = Ico a b ↔ ¬a < b where + mp h := by simpa using Set.ext_iff.mp h a + mpr h := by rw [Ioo_eq_empty h, Ico_eq_empty h] + +-- Mirrored versions of the above for `simp`. + +@[simp] theorem Ioc_eq_Icc_same_iff : Ioc a b = Icc a b ↔ ¬a ≤ b := + eq_comm.trans Icc_eq_Ioc_same_iff + +@[simp] theorem Ico_eq_Icc_same_iff : Ico a b = Icc a b ↔ ¬a ≤ b := + eq_comm.trans Icc_eq_Ico_same_iff + +@[simp] theorem Ioo_eq_Icc_same_iff : Ioo a b = Icc a b ↔ ¬a ≤ b := + eq_comm.trans Icc_eq_Ioo_same_iff + +@[simp] theorem Ico_eq_Ioc_same_iff : Ico a b = Ioc a b ↔ ¬a < b := + eq_comm.trans Ioc_eq_Ico_same_iff + +@[simp] theorem Ioc_eq_Ioo_same_iff : Ioc a b = Ioo a b ↔ ¬a < b := + eq_comm.trans Ioo_eq_Ioc_same_iff + +@[simp] theorem Ico_eq_Ioo_same_iff : Ico a b = Ioo a b ↔ ¬a < b := + eq_comm.trans Ioo_eq_Ico_same_iff + +end matched_intervals + end Preorder section PartialOrder @@ -1460,6 +1508,7 @@ theorem Ioc_union_Ioc_union_Ioc_cycle : all_goals solve_by_elim (config := { maxDepth := 5 }) [min_le_of_left_le, min_le_of_right_le, le_max_of_le_left, le_max_of_le_right, le_refl] + end LinearOrder /-! diff --git a/Mathlib/Topology/Order/Compact.lean b/Mathlib/Topology/Order/Compact.lean index 865a0c4150338..041ed9ae203b1 100644 --- a/Mathlib/Topology/Order/Compact.lean +++ b/Mathlib/Topology/Order/Compact.lean @@ -155,6 +155,26 @@ instance compactSpace_Icc (a b : α) : CompactSpace (Icc a b) := end +section openIntervals +variable {α : Type*} [LinearOrder α] [TopologicalSpace α] [OrderTopology α] [DenselyOrdered α] + +/-- `Set.Ico a b` is only compact if it is empty. -/ +@[simp] +theorem isCompact_Ico_iff {a b : α} : IsCompact (Set.Ico a b) ↔ b ≤ a := + ⟨fun h => isClosed_Ico_iff.mp h.isClosed, by simp_all⟩ + +/-- `Set.Ioc a b` is only compact if it is empty. -/ +@[simp] +theorem isCompact_Ioc_iff {a b : α} : IsCompact (Set.Ioc a b) ↔ b ≤ a := + ⟨fun h => isClosed_Ioc_iff.mp h.isClosed, by simp_all⟩ + +/-- `Set.Ioo a b` is only compact if it is empty. -/ +@[simp] +theorem isCompact_Ioo_iff {a b : α} : IsCompact (Set.Ioo a b) ↔ b ≤ a := + ⟨fun h => isClosed_Ioo_iff.mp h.isClosed, by simp_all⟩ + +end openIntervals + /-! ### Extreme value theorem -/ diff --git a/Mathlib/Topology/Order/DenselyOrdered.lean b/Mathlib/Topology/Order/DenselyOrdered.lean index f7a40b1ae2311..0c981229bfc16 100644 --- a/Mathlib/Topology/Order/DenselyOrdered.lean +++ b/Mathlib/Topology/Order/DenselyOrdered.lean @@ -374,8 +374,7 @@ theorem Dense.exists_countable_dense_subset_no_bot_top [Nontrivial α] {s : Set · simp [hx] · simp [hx] -variable (α) - +variable (α) in /-- If `α` is a nontrivial separable dense linear order, then there exists a countable dense set `s : Set α` that contains neither top nor bottom elements of `α`. For a dense set containing both bot and top elements, see @@ -384,4 +383,28 @@ theorem exists_countable_dense_no_bot_top [SeparableSpace α] [Nontrivial α] : ∃ s : Set α, s.Countable ∧ Dense s ∧ (∀ x, IsBot x → x ∉ s) ∧ ∀ x, IsTop x → x ∉ s := by simpa using dense_univ.exists_countable_dense_subset_no_bot_top +/-- `Set.Ico a b` is only closed if it is empty. -/ +@[simp] +theorem isClosed_Ico_iff {a b : α} : IsClosed (Set.Ico a b) ↔ b ≤ a := by + refine ⟨fun h => le_of_not_lt fun hab => ?_, by simp_all⟩ + have := h.closure_eq + rw [closure_Ico hab.ne, Icc_eq_Ico_same_iff] at this + exact this hab.le + +/-- `Set.Ioc a b` is only closed if it is empty. -/ +@[simp] +theorem isClosed_Ioc_iff {a b : α} : IsClosed (Set.Ioc a b) ↔ b ≤ a := by + refine ⟨fun h => le_of_not_lt fun hab => ?_, by simp_all⟩ + have := h.closure_eq + rw [closure_Ioc hab.ne, Icc_eq_Ioc_same_iff] at this + exact this hab.le + +/-- `Set.Ioo a b` is only closed if it is empty. -/ +@[simp] +theorem isClosed_Ioo_iff {a b : α} : IsClosed (Set.Ioo a b) ↔ b ≤ a := by + refine ⟨fun h => le_of_not_lt fun hab => ?_, by simp_all⟩ + have := h.closure_eq + rw [closure_Ioo hab.ne, Icc_eq_Ioo_same_iff] at this + exact this hab.le + end DenselyOrdered From bbc7c116b250404339cb2a529a9e641d5e2d510b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 5 Jan 2025 13:11:08 +0000 Subject: [PATCH 163/189] feat: if `f` is a measurable group hom, then every point has a neighborhood `s` such that `f '' s` is bounded (#20304) This will be used to solve Cauchy's functional equation. From LeanCamCombi Co-authored-by: Mantas Baksys --- Mathlib/Data/Set/Pointwise/SMul.lean | 5 ++ Mathlib/MeasureTheory/Measure/Haar/Basic.lean | 63 ++++++++++++------- .../Measure/Haar/NormedSpace.lean | 33 ++++++++-- 3 files changed, 75 insertions(+), 26 deletions(-) diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index 4a60e0e49e98f..84372884b6f19 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -202,6 +202,11 @@ theorem image_smul_distrib [MulOneClass α] [MulOneClass β] [FunLike F α β] [ f '' (a • s) = f a • f '' s := image_comm <| map_mul _ _ +open scoped RightActions in +@[to_additive] +lemma image_op_smul_distrib [MulOneClass α] [MulOneClass β] [FunLike F α β] [MonoidHomClass F α β] + (f : F) (a : α) (s : Set α) : f '' (s <• a) = f '' s <• f a := image_comm fun _ ↦ map_mul _ _ _ + section SMul variable [SMul αᵐᵒᵖ β] [SMul β γ] [SMul α γ] diff --git a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean index caed697da9d50..65acdbcc140ed 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean @@ -590,33 +590,52 @@ inner regular Haar measures (and in particular for any Haar measure on a second open Pointwise -/-- **Steinhaus Theorem** In any locally compact group `G` with an inner regular Haar measure `μ`, +@[to_additive] +private lemma steinhaus_mul_aux (μ : Measure G) [IsHaarMeasure μ] [μ.InnerRegularCompactLTTop] + [LocallyCompactSpace G] (E : Set G) (hE : MeasurableSet E) + (hEapprox : ∃ K ⊆ E, IsCompact K ∧ 0 < μ K) : E / E ∈ 𝓝 (1 : G) := by + /- For any measure `μ` and set `E` containing a compact set `K` of positive measure, there exists + a neighborhood `V` of the identity such that `v • K \ K` has small measure for all `v ∈ V`, say + `< μ K`. Then `v • K` and `K` can not be disjoint, as otherwise `μ (v • K \ K) = μ (v • K) = μ K`. + This show that `K / K` contains the neighborhood `V` of `1`, and therefore that it is + itself such a neighborhood. -/ + obtain ⟨K, hKE, hK, K_closed, hKpos⟩ : ∃ K ⊆ E, IsCompact K ∧ IsClosed K ∧ 0 < μ K := by + obtain ⟨K, hKE, hK_comp, hK_meas⟩ := hEapprox + exact ⟨closure K, hK_comp.closure_subset_measurableSet hE hKE, hK_comp.closure, + isClosed_closure, by rwa [hK_comp.measure_closure]⟩ + filter_upwards [eventually_nhds_one_measure_smul_diff_lt hK K_closed hKpos.ne' (μ := μ)] with g hg + obtain ⟨_, ⟨x, hxK, rfl⟩, hgxK⟩ : ∃ x ∈ g • K, x ∈ K := + not_disjoint_iff.1 fun hd ↦ by simp [hd.symm.sdiff_eq_right, measure_smul] at hg + simpa using div_mem_div (hKE hgxK) (hKE hxK) + +/-- **Steinhaus Theorem** for finite mass sets. + +In any locally compact group `G` with an Haar measure `μ` that's inner regular on finite measure +sets, for any measurable set `E` of finite positive measure, the set `E / E` is a neighbourhood of +`1`. -/ +@[to_additive +"**Steinhaus Theorem** for finite mass sets. + +In any locally compact group `G` with an Haar measure `μ` that's inner regular on finite measure +sets, for any measurable set `E` of finite positive measure, the set `E - E` is a neighbourhood of +`0`. "] +theorem div_mem_nhds_one_of_haar_pos_ne_top (μ : Measure G) [IsHaarMeasure μ] + [LocallyCompactSpace G] [μ.InnerRegularCompactLTTop] (E : Set G) (hE : MeasurableSet E) + (hEpos : 0 < μ E) (hEfin : μ E ≠ ∞) : E / E ∈ 𝓝 (1 : G) := + steinhaus_mul_aux μ E hE <| hE.exists_lt_isCompact_of_ne_top hEfin hEpos + +/-- **Steinhaus Theorem**. + +In any locally compact group `G` with an inner regular Haar measure `μ`, for any measurable set `E` of positive measure, the set `E / E` is a neighbourhood of `1`. -/ @[to_additive -"**Steinhaus Theorem** In any locally compact group `G` with an inner regular Haar measure `μ`, +"**Steinhaus Theorem**. + +In any locally compact group `G` with an inner regular Haar measure `μ`, for any measurable set `E` of positive measure, the set `E - E` is a neighbourhood of `0`."] theorem div_mem_nhds_one_of_haar_pos (μ : Measure G) [IsHaarMeasure μ] [LocallyCompactSpace G] [InnerRegular μ] (E : Set G) (hE : MeasurableSet E) (hEpos : 0 < μ E) : - E / E ∈ 𝓝 (1 : G) := by - /- For any inner regular measure `μ` and set `E` of positive measure, we can find a compact - set `K` of positive measure inside `E`. Further, there exists a neighborhood `V` of the - identity such that `v • K \ K` has small measure for all `v ∈ V`, say `< μ K`. - Then `v • K` and `K` can not be disjoint, as otherwise `μ (v • K \ K) = μ (v • K) = μ K`. - This show that `K / K` contains the neighborhood `V` of `1`, and therefore that it is - itself such a neighborhood. -/ - obtain ⟨K, hKE, hK, K_closed, hKpos⟩ : - ∃ (K : Set G), K ⊆ E ∧ IsCompact K ∧ IsClosed K ∧ 0 < μ K := by - rcases MeasurableSet.exists_lt_isCompact hE hEpos with ⟨K, KE, K_comp, K_meas⟩ - refine ⟨closure K, ?_, K_comp.closure, isClosed_closure, ?_⟩ - · exact K_comp.closure_subset_measurableSet hE KE - · rwa [K_comp.measure_closure] - filter_upwards [eventually_nhds_one_measure_smul_diff_lt hK K_closed hKpos.ne' (μ := μ)] with g hg - have : ¬Disjoint (g • K) K := fun hd ↦ by - rw [hd.symm.sdiff_eq_right, measure_smul] at hg - exact hg.false - rcases Set.not_disjoint_iff.1 this with ⟨_, ⟨x, hxK, rfl⟩, hgxK⟩ - simpa using div_mem_div (hKE hgxK) (hKE hxK) - + E / E ∈ 𝓝 (1 : G) := steinhaus_mul_aux μ E hE <| hE.exists_lt_isCompact hEpos section SecondCountable_SigmaFinite /-! In this section, we investigate uniqueness of left-invariant measures without assuming that diff --git a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean index dc9e9ec5eeea0..c52acf767b115 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean @@ -14,12 +14,9 @@ import Mathlib.MeasureTheory.Integral.SetIntegral noncomputable section +open Function Filter Inv MeasureTheory.Measure Module Set TopologicalSpace open scoped NNReal ENNReal Pointwise Topology -open Inv Set Function MeasureTheory.Measure Filter - -open Module - namespace MeasureTheory namespace Measure @@ -44,6 +41,34 @@ instance MapLinearEquiv.isAddHaarMeasure (e : G ≃ₗ[𝕜] H) : IsAddHaarMeasu end LinearEquiv +section SeminormedGroup +variable {G H : Type*} [MeasurableSpace G] [Group G] [TopologicalSpace G] + [TopologicalGroup G] [BorelSpace G] [LocallyCompactSpace G] + [MeasurableSpace H] [SeminormedGroup H] [OpensMeasurableSpace H] + +-- TODO: This could be streamlined by proving that inner regular always exist +open Metric Bornology in +@[to_additive] +lemma _root_.MonoidHom.exists_nhds_isBounded (f : G →* H) (hf : Measurable f) (x : G) : + ∃ s ∈ 𝓝 x, IsBounded (f '' s) := by + let K : PositiveCompacts G := Classical.arbitrary _ + obtain ⟨n, hn⟩ : ∃ n : ℕ, 0 < haar (interior K ∩ f ⁻¹' ball 1 n) := by + by_contra! + simp_rw [nonpos_iff_eq_zero, ← measure_iUnion_null_iff, ← inter_iUnion, ← preimage_iUnion, + iUnion_ball_nat, preimage_univ, inter_univ] at this + exact this.not_gt <| isOpen_interior.measure_pos _ K.interior_nonempty + rw [← one_mul x, ← op_smul_eq_mul] + refine ⟨_, smul_mem_nhds_smul _ <| div_mem_nhds_one_of_haar_pos_ne_top haar _ + (isOpen_interior.measurableSet.inter <| hf measurableSet_ball) hn <| + mt (measure_mono_top <| inter_subset_left.trans interior_subset) K.isCompact.measure_ne_top, + ?_⟩ + have : Bornology.IsBounded (f '' (interior K ∩ f ⁻¹' ball 1 n)) := + isBounded_ball.subset <| (image_mono inter_subset_right).trans <| image_preimage_subset _ _ + rw [image_op_smul_distrib, image_div] + exact (this.div this).smul _ + +end SeminormedGroup + variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [MeasurableSpace E] [BorelSpace E] [FiniteDimensional ℝ E] (μ : Measure E) [IsAddHaarMeasure μ] {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] From 599c46e994f3d200d5b4ffaeeea9ab89de8d7d46 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Sun, 5 Jan 2025 13:11:09 +0000 Subject: [PATCH 164/189] feat: a conditional kernel is almost everywhere a probability measure (#20430) --- .../Function/AEEqOfLIntegral.lean | 5 ++ .../Kernel/Disintegration/Basic.lean | 58 ++++++++++++++++++- 2 files changed, 61 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Function/AEEqOfLIntegral.lean b/Mathlib/MeasureTheory/Function/AEEqOfLIntegral.lean index 6e2b5cedfb464..41ac33f231c08 100644 --- a/Mathlib/MeasureTheory/Function/AEEqOfLIntegral.lean +++ b/Mathlib/MeasureTheory/Function/AEEqOfLIntegral.lean @@ -69,6 +69,11 @@ theorem ae_const_le_iff_forall_lt_measure_zero {β} [LinearOrder β] [Topologica intro n exact hc _ (u_lt n) +lemma ae_le_const_iff_forall_gt_measure_zero {β} [LinearOrder β] [TopologicalSpace β] + [OrderTopology β] [FirstCountableTopology β] {μ : Measure α} (f : α → β) (c : β) : + (∀ᵐ x ∂μ, f x ≤ c) ↔ ∀ b, c < b → μ {x | b ≤ f x} = 0 := + ae_const_le_iff_forall_lt_measure_zero (β := βᵒᵈ) _ _ + theorem ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ [SigmaFinite μ] {f g : α → ℝ≥0∞} (hf : AEMeasurable f μ) (h : ∀ s, MeasurableSet s → μ s < ∞ → ∫⁻ x in s, f x ∂μ ≤ ∫⁻ x in s, g x ∂μ) : diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 8d443a75e791c..aeae2b78f63d7 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Yaël Dillies, Kin Yau James Wong. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Kin Yau James Wong, Rémy Degenne -/ +import Mathlib.MeasureTheory.Function.AEEqOfLIntegral import Mathlib.Probability.Kernel.Composition.MeasureCompProd /-! @@ -126,9 +127,62 @@ class IsCondKernel : Prop where instance instIsCondKernel_zero (κCond : Kernel (α × β) Ω) : IsCondKernel 0 κCond where disintegrate := by simp -variable [κ.IsCondKernel κCond] +lemma disintegrate [κ.IsCondKernel κCond] : κ.fst ⊗ₖ κCond = κ := IsCondKernel.disintegrate + +/-- A conditional kernel is almost everywhere a probability measure. -/ +lemma IsCondKernel.isProbabilityMeasure_ae [IsFiniteKernel κ.fst] [κ.IsCondKernel κCond] (a : α) : + ∀ᵐ b ∂(κ.fst a), IsProbabilityMeasure (κCond (a, b)) := by + have h := disintegrate κ κCond + by_cases h_sfin : IsSFiniteKernel κCond + swap; · rw [Kernel.compProd_of_not_isSFiniteKernel_right _ _ h_sfin] at h; simp [h.symm] + suffices ∀ᵐ b ∂(κ.fst a), κCond (a, b) Set.univ = 1 by + convert this with b + exact ⟨fun _ ↦ measure_univ, fun h ↦ ⟨h⟩⟩ + suffices (∀ᵐ b ∂(κ.fst a), κCond (a, b) Set.univ ≤ 1) + ∧ (∀ᵐ b ∂(κ.fst a), 1 ≤ κCond (a, b) Set.univ) by + filter_upwards [this.1, this.2] with b h1 h2 using le_antisymm h1 h2 + have h_eq s (hs : MeasurableSet s) : + ∫⁻ b, s.indicator (fun b ↦ κCond (a, b) Set.univ) b ∂κ.fst a = κ.fst a s := by + conv_rhs => rw [← h] + rw [fst_compProd_apply _ _ _ hs] + have h_meas : Measurable fun b ↦ κCond (a, b) Set.univ := + (κCond.measurable_coe MeasurableSet.univ).comp measurable_prod_mk_left + constructor + · rw [ae_le_const_iff_forall_gt_measure_zero] + intro r hr + let s := {b | r ≤ κCond (a, b) Set.univ} + have hs : MeasurableSet s := h_meas measurableSet_Ici + have h_2_le : s.indicator (fun _ ↦ r) ≤ s.indicator (fun b ↦ (κCond (a, b)) Set.univ) := by + intro b + by_cases hbs : b ∈ s + · simpa [hbs] + · simp [hbs] + have : ∫⁻ b, s.indicator (fun _ ↦ r) b ∂(κ.fst a) ≤ κ.fst a s := + (lintegral_mono h_2_le).trans_eq (h_eq s hs) + rw [lintegral_indicator_const hs] at this + by_contra h_ne_zero + rw [← not_lt] at this + refine this ?_ + conv_lhs => rw [← one_mul (κ.fst a s)] + exact ENNReal.mul_lt_mul_right' h_ne_zero (measure_ne_top _ _) hr + · rw [ae_const_le_iff_forall_lt_measure_zero] + intro r hr + let s := {b | κCond (a, b) Set.univ ≤ r} + have hs : MeasurableSet s := h_meas measurableSet_Iic + have h_2_le : s.indicator (fun b ↦ (κCond (a, b)) Set.univ) ≤ s.indicator (fun _ ↦ r) := by + intro b + by_cases hbs : b ∈ s + · simpa [hbs] + · simp [hbs] + have : κ.fst a s ≤ ∫⁻ b, s.indicator (fun _ ↦ r) b ∂(κ.fst a) := + (h_eq s hs).symm.trans_le (lintegral_mono h_2_le) + rw [lintegral_indicator_const hs] at this + by_contra h_ne_zero + rw [← not_lt] at this + refine this ?_ + conv_rhs => rw [← one_mul (κ.fst a s)] + exact ENNReal.mul_lt_mul_right' h_ne_zero (measure_ne_top _ _) hr -lemma disintegrate : κ.fst ⊗ₖ κCond = κ := IsCondKernel.disintegrate /-! #### Existence of a disintegrating kernel in a countable space -/ From 9173268a35c4ba73d7a0315edfc5233a13c1e053 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 5 Jan 2025 13:11:10 +0000 Subject: [PATCH 165/189] feat(Asymptotics): add `Asymptotics.*.*Prod` lemmas (#20458) `IsBigO.fisnetProd` comes from https://github.com/urkud/SardMoreira, the others are added to avoid gaps in the API. --- .../Asymptotics/AsymptoticEquivalent.lean | 25 +++++-- Mathlib/Analysis/Asymptotics/Asymptotics.lean | 66 ++++++++++++++++--- Mathlib/Analysis/Asymptotics/Theta.lean | 14 ++++ 3 files changed, 91 insertions(+), 14 deletions(-) diff --git a/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean b/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean index d0a29face182e..ef49406727787 100644 --- a/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean +++ b/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean @@ -254,12 +254,29 @@ end SMul section mul_inv -variable {α β : Type*} [NormedField β] {t u v w : α → β} {l : Filter α} +variable {α ι β : Type*} [NormedField β] {t u v w : α → β} {l : Filter α} -theorem IsEquivalent.mul (htu : t ~[l] u) (hvw : v ~[l] w) : t * v ~[l] u * w := +protected theorem IsEquivalent.mul (htu : t ~[l] u) (hvw : v ~[l] w) : t * v ~[l] u * w := htu.smul hvw -theorem IsEquivalent.inv (huv : u ~[l] v) : (fun x ↦ (u x)⁻¹) ~[l] fun x ↦ (v x)⁻¹ := by +theorem IsEquivalent.listProd {L : List ι} {f g : ι → α → β} (h : ∀ i ∈ L, f i ~[l] g i) : + (fun x ↦ (L.map (f · x)).prod) ~[l] (fun x ↦ (L.map (g · x)).prod) := by + induction L with + | nil => simp [IsEquivalent.refl] + | cons i L ihL => + simp only [List.forall_mem_cons, List.map_cons, List.prod_cons] at h ⊢ + exact h.1.mul (ihL h.2) + +theorem IsEquivalent.multisetProd {s : Multiset ι} {f g : ι → α → β} (h : ∀ i ∈ s, f i ~[l] g i) : + (fun x ↦ (s.map (f · x)).prod) ~[l] (fun x ↦ (s.map (g · x)).prod) := by + obtain ⟨l, rfl⟩ : ∃ l : List ι, ↑l = s := Quotient.mk_surjective s + exact listProd h + +theorem IsEquivalent.finsetProd {s : Finset ι} {f g : ι → α → β} (h : ∀ i ∈ s, f i ~[l] g i) : + (∏ i ∈ s, f i ·) ~[l] (∏ i ∈ s, g i ·) := + multisetProd h + +protected theorem IsEquivalent.inv (huv : u ~[l] v) : (fun x ↦ (u x)⁻¹) ~[l] fun x ↦ (v x)⁻¹ := by rw [isEquivalent_iff_exists_eq_mul] at * rcases huv with ⟨φ, hφ, h⟩ rw [← inv_one] @@ -267,7 +284,7 @@ theorem IsEquivalent.inv (huv : u ~[l] v) : (fun x ↦ (u x)⁻¹) ~[l] fun x convert h.inv simp [mul_comm] -theorem IsEquivalent.div (htu : t ~[l] u) (hvw : v ~[l] w) : +protected theorem IsEquivalent.div (htu : t ~[l] u) (hvw : v ~[l] w) : (fun x ↦ t x / v x) ~[l] fun x ↦ u x / w x := by simpa only [div_eq_mul_inv] using htu.mul hvw.inv diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index 74755cbeef3a1..b38690a42c515 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -1570,11 +1570,11 @@ variable {ι : Type*} {A : ι → α → E'} {C : ι → ℝ} {s : Finset ι} theorem IsBigOWith.sum (h : ∀ i ∈ s, IsBigOWith (C i) l (A i) g) : IsBigOWith (∑ i ∈ s, C i) l (fun x => ∑ i ∈ s, A i x) g := by - classical - induction' s using Finset.induction_on with i s is IH - · simp only [isBigOWith_zero', Finset.sum_empty, forall_true_iff] - · simp only [is, Finset.sum_insert, not_false_iff] - exact (h _ (Finset.mem_insert_self i s)).add (IH fun j hj => h _ (Finset.mem_insert_of_mem hj)) + induction s using Finset.cons_induction with + | empty => simp only [isBigOWith_zero', Finset.sum_empty, forall_true_iff] + | cons i s is IH => + simp only [is, Finset.sum_cons, Finset.forall_mem_cons] at h ⊢ + exact h.1.add (IH h.2) theorem IsBigO.sum (h : ∀ i ∈ s, A i =O[l] g) : (fun x => ∑ i ∈ s, A i x) =O[l] g := by simp only [IsBigO_def] at * @@ -1582,14 +1582,60 @@ theorem IsBigO.sum (h : ∀ i ∈ s, A i =O[l] g) : (fun x => ∑ i ∈ s, A i x exact ⟨_, IsBigOWith.sum hC⟩ theorem IsLittleO.sum (h : ∀ i ∈ s, A i =o[l] g') : (fun x => ∑ i ∈ s, A i x) =o[l] g' := by - classical - induction' s using Finset.induction_on with i s is IH - · simp only [isLittleO_zero, Finset.sum_empty, forall_true_iff] - · simp only [is, Finset.sum_insert, not_false_iff] - exact (h _ (Finset.mem_insert_self i s)).add (IH fun j hj => h _ (Finset.mem_insert_of_mem hj)) + simp only [← Finset.sum_apply] + exact Finset.sum_induction A (· =o[l] g') (fun _ _ ↦ .add) (isLittleO_zero ..) h end Sum +section Prod +variable {ι : Type*} + +theorem IsBigO.listProd {L : List ι} {f : ι → α → R} {g : ι → α → 𝕜} + (hf : ∀ i ∈ L, f i =O[l] g i) : + (fun x ↦ (L.map (f · x)).prod) =O[l] (fun x ↦ (L.map (g · x)).prod) := by + induction L with + | nil => simp [isBoundedUnder_const] + | cons i L ihL => + simp only [List.map_cons, List.prod_cons, List.forall_mem_cons] at hf ⊢ + exact hf.1.mul (ihL hf.2) + +theorem IsBigO.multisetProd {R 𝕜 : Type*} [SeminormedCommRing R] [NormedField 𝕜] + {s : Multiset ι} {f : ι → α → R} {g : ι → α → 𝕜} (hf : ∀ i ∈ s, f i =O[l] g i) : + (fun x ↦ (s.map (f · x)).prod) =O[l] (fun x ↦ (s.map (g · x)).prod) := by + obtain ⟨l, rfl⟩ : ∃ l : List ι, ↑l = s := Quotient.mk_surjective s + exact mod_cast IsBigO.listProd hf + +theorem IsBigO.finsetProd {R 𝕜 : Type*} [SeminormedCommRing R] [NormedField 𝕜] + {s : Finset ι} {f : ι → α → R} {g : ι → α → 𝕜} + (hf : ∀ i ∈ s, f i =O[l] g i) : (∏ i ∈ s, f i ·) =O[l] (∏ i ∈ s, g i ·) := + .multisetProd hf + +theorem IsLittleO.listProd {L : List ι} {f : ι → α → R} {g : ι → α → 𝕜} + (h₁ : ∀ i ∈ L, f i =O[l] g i) (h₂ : ∃ i ∈ L, f i =o[l] g i) : + (fun x ↦ (L.map (f · x)).prod) =o[l] (fun x ↦ (L.map (g · x)).prod) := by + induction L with + | nil => simp at h₂ + | cons i L ihL => + simp only [List.map_cons, List.prod_cons, List.forall_mem_cons, List.exists_mem_cons_iff] + at h₁ h₂ ⊢ + cases h₂ with + | inl hi => exact hi.mul_isBigO <| .listProd h₁.2 + | inr hL => exact h₁.1.mul_isLittleO <| ihL h₁.2 hL + +theorem IsLittleO.multisetProd {R 𝕜 : Type*} [SeminormedCommRing R] [NormedField 𝕜] + {s : Multiset ι} {f : ι → α → R} {g : ι → α → 𝕜} (h₁ : ∀ i ∈ s, f i =O[l] g i) + (h₂ : ∃ i ∈ s, f i =o[l] g i) : + (fun x ↦ (s.map (f · x)).prod) =o[l] (fun x ↦ (s.map (g · x)).prod) := by + obtain ⟨l, rfl⟩ : ∃ l : List ι, ↑l = s := Quotient.mk_surjective s + exact mod_cast IsLittleO.listProd h₁ h₂ + +theorem IsLittleO.finsetProd {R 𝕜 : Type*} [SeminormedCommRing R] [NormedField 𝕜] + {s : Finset ι} {f : ι → α → R} {g : ι → α → 𝕜} (h₁ : ∀ i ∈ s, f i =O[l] g i) + (h₂ : ∃ i ∈ s, f i =o[l] g i) : (∏ i ∈ s, f i ·) =o[l] (∏ i ∈ s, g i ·) := + .multisetProd h₁ h₂ + +end Prod + /-! ### Relation between `f = o(g)` and `f / g → 0` -/ diff --git a/Mathlib/Analysis/Asymptotics/Theta.lean b/Mathlib/Analysis/Asymptotics/Theta.lean index f41f4161c7d45..2ec951024df70 100644 --- a/Mathlib/Analysis/Asymptotics/Theta.lean +++ b/Mathlib/Analysis/Asymptotics/Theta.lean @@ -204,6 +204,20 @@ theorem IsTheta.mul {f₁ f₂ : α → 𝕜} {g₁ g₂ : α → 𝕜'} (h₁ : (fun x ↦ f₁ x * f₂ x) =Θ[l] fun x ↦ g₁ x * g₂ x := h₁.smul h₂ +theorem IsTheta.listProd {ι : Type*} {L : List ι} {f : ι → α → 𝕜} {g : ι → α → 𝕜'} + (h : ∀ i ∈ L, f i =Θ[l] g i) : + (fun x ↦ (L.map (f · x)).prod) =Θ[l] (fun x ↦ (L.map (g · x)).prod) := + ⟨.listProd fun i hi ↦ (h i hi).isBigO, .listProd fun i hi ↦ (h i hi).symm.isBigO⟩ + +theorem IsTheta.multisetProd {ι : Type*} {s : Multiset ι} {f : ι → α → 𝕜} {g : ι → α → 𝕜'} + (h : ∀ i ∈ s, f i =Θ[l] g i) : + (fun x ↦ (s.map (f · x)).prod) =Θ[l] (fun x ↦ (s.map (g · x)).prod) := + ⟨.multisetProd fun i hi ↦ (h i hi).isBigO, .multisetProd fun i hi ↦ (h i hi).symm.isBigO⟩ + +theorem IsTheta.finsetProd {ι : Type*} {s : Finset ι} {f : ι → α → 𝕜} {g : ι → α → 𝕜'} + (h : ∀ i ∈ s, f i =Θ[l] g i) : (∏ i ∈ s, f i ·) =Θ[l] (∏ i ∈ s, g i ·) := + ⟨.finsetProd fun i hi ↦ (h i hi).isBigO, .finsetProd fun i hi ↦ (h i hi).symm.isBigO⟩ + theorem IsTheta.inv {f : α → 𝕜} {g : α → 𝕜'} (h : f =Θ[l] g) : (fun x ↦ (f x)⁻¹) =Θ[l] fun x ↦ (g x)⁻¹ := ⟨h.2.inv_rev h.1.eq_zero_imp, h.1.inv_rev h.2.eq_zero_imp⟩ From 2590de306b6e8f6b8cbc94b07f0e2d50d282efcd Mon Sep 17 00:00:00 2001 From: smorel394 Date: Sun, 5 Jan 2025 13:43:37 +0000 Subject: [PATCH 166/189] chore(CategoryTheory/Adjunction.Basic, CategoryTheory/Equivalence): change definition of `Adjunction.comp` and `Equivalence.trans` (#20490) The definitions of `Adjunction.comp` and `Equivalence.trans` were making abuse of `defeq`. We fix this in a way that makes `Equivalence.trans_toAdjunction` true by `rfl`. We also add a lemma `Equivalence.refl_toAdjunction`. --- .../DoldKan/Compatibility.lean | 8 ++++---- Mathlib/CategoryTheory/Adjunction/Basic.lean | 14 ++++++++------ Mathlib/CategoryTheory/Adjunction/Mates.lean | 5 ++--- Mathlib/CategoryTheory/Equivalence.lean | 17 ++++++++--------- .../CategoryTheory/Localization/Predicate.lean | 4 +++- 5 files changed, 25 insertions(+), 23 deletions(-) diff --git a/Mathlib/AlgebraicTopology/DoldKan/Compatibility.lean b/Mathlib/AlgebraicTopology/DoldKan/Compatibility.lean index c8235ee654872..cf5fcca453bd2 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Compatibility.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Compatibility.lean @@ -122,8 +122,8 @@ theorem equivalence₂CounitIso_eq : (equivalence₂ eB hF).counitIso = equivalence₂CounitIso eB hF := by ext Y' dsimp [equivalence₂, Iso.refl] - simp only [equivalence₁CounitIso_eq, equivalence₂CounitIso_hom_app, - equivalence₁CounitIso_hom_app, Functor.map_comp, assoc] + simp only [equivalence₁CounitIso_eq, equivalence₁CounitIso_hom_app, comp_id, id_comp, + Functor.map_comp, assoc, equivalence₂CounitIso_hom_app] /-- The unit isomorphism of the equivalence `equivalence₂` between `A` and `B`. -/ @[simps!] @@ -138,8 +138,8 @@ def equivalence₂UnitIso : 𝟭 A ≅ (F ⋙ eB.inverse) ⋙ eB.functor ⋙ e'. theorem equivalence₂UnitIso_eq : (equivalence₂ eB hF).unitIso = equivalence₂UnitIso eB hF := by ext X dsimp [equivalence₂] - simp only [equivalence₂UnitIso_hom_app, equivalence₁UnitIso_eq, equivalence₁UnitIso_hom_app, - assoc, NatIso.cancel_natIso_hom_left] + simp only [equivalence₁UnitIso_eq, equivalence₁UnitIso_hom_app, comp_id, id_comp, assoc, + equivalence₂UnitIso_hom_app] rfl variable {eB} diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index 9ba41bce3a283..49a73af0c7d52 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -516,12 +516,14 @@ variable {E : Type u₃} [ℰ : Category.{v₃} E] {H : D ⥤ E} {I : E ⥤ D} See . -/ +@[simps! (config := .lemmasOnly) unit counit] def comp : F ⋙ H ⊣ I ⋙ G := mk' { homEquiv := fun _ _ ↦ Equiv.trans (adj₂.homEquiv _ _) (adj₁.homEquiv _ _) - unit := adj₁.unit ≫ (whiskerLeft F <| whiskerRight adj₂.unit G) ≫ (Functor.associator _ _ _).inv - counit := - (Functor.associator _ _ _).hom ≫ (whiskerLeft I <| whiskerRight adj₁.counit H) ≫ adj₂.counit } + unit := adj₁.unit ≫ whiskerRight (F.rightUnitor.inv ≫ whiskerLeft F adj₂.unit ≫ + (Functor.associator _ _ _ ).inv) G ≫ (Functor.associator _ _ _).hom + counit := (Functor.associator _ _ _ ).inv ≫ whiskerRight ((Functor.associator _ _ _ ).hom ≫ + whiskerLeft _ adj₁.counit ≫ I.rightUnitor.hom) _ ≫ adj₂.counit } @[simp, reassoc] lemma comp_unit_app (X : C) : @@ -668,10 +670,10 @@ lemma isLeftAdjoint_inverse : e.inverse.IsLeftAdjoint := lemma isRightAdjoint_functor : e.functor.IsRightAdjoint := e.symm.isRightAdjoint_inverse +lemma refl_toAdjunction : (refl (C := C)).toAdjunction = Adjunction.id := rfl + lemma trans_toAdjunction {E : Type*} [Category E] (e' : D ≌ E) : - (e.trans e').toAdjunction = e.toAdjunction.comp e'.toAdjunction := by - ext - simp [trans] + (e.trans e').toAdjunction = e.toAdjunction.comp e'.toAdjunction := rfl end Equivalence diff --git a/Mathlib/CategoryTheory/Adjunction/Mates.lean b/Mathlib/CategoryTheory/Adjunction/Mates.lean index ca344b3aa3796..cd608363c1820 100644 --- a/Mathlib/CategoryTheory/Adjunction/Mates.lean +++ b/Mathlib/CategoryTheory/Adjunction/Mates.lean @@ -208,9 +208,8 @@ theorem mateEquiv_hcomp rightAdjointSquare.hcomp (mateEquiv adj₁ adj₂ α) (mateEquiv adj₃ adj₄ β) := by unfold leftAdjointSquare.hcomp rightAdjointSquare.hcomp mateEquiv Adjunction.comp ext c - simp only [comp_obj, mk'_unit, whiskerLeft_comp, whiskerLeft_twice, mk'_counit, whiskerRight_comp, - assoc, Equiv.coe_fn_mk, comp_app, whiskerLeft_app, whiskerRight_app, id_obj, associator_inv_app, - Functor.comp_map, associator_hom_app, map_id, id_comp, whiskerRight_twice] + dsimp + simp only [comp_id, map_comp, id_comp, assoc] slice_rhs 2 4 => rw [← R₂.map_comp, ← R₂.map_comp, ← assoc, ← unit_naturality (adj₄)] rw [R₂.map_comp, L₄.map_comp, R₄.map_comp, R₂.map_comp] diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean index f6f45ef4bd954..949a01797a9d6 100644 --- a/Mathlib/CategoryTheory/Equivalence.lean +++ b/Mathlib/CategoryTheory/Equivalence.lean @@ -269,20 +269,19 @@ variable {E : Type u₃} [Category.{v₃} E] def trans (e : C ≌ D) (f : D ≌ E) : C ≌ E where functor := e.functor ⋙ f.functor inverse := f.inverse ⋙ e.inverse - unitIso := by - refine Iso.trans e.unitIso ?_ - exact isoWhiskerLeft e.functor (isoWhiskerRight f.unitIso e.inverse) - counitIso := by - refine Iso.trans ?_ f.counitIso - exact isoWhiskerLeft f.inverse (isoWhiskerRight e.counitIso f.functor) + unitIso := e.unitIso ≪≫ isoWhiskerRight (e.functor.rightUnitor.symm ≪≫ + isoWhiskerLeft _ f.unitIso ≪≫ (Functor.associator _ _ _ ).symm) _ ≪≫ Functor.associator _ _ _ + counitIso := (Functor.associator _ _ _ ).symm ≪≫ isoWhiskerRight ((Functor.associator _ _ _ ) ≪≫ + isoWhiskerLeft _ e.counitIso ≪≫ f.inverse.rightUnitor) _ ≪≫ f.counitIso -- We wouldn't have needed to give this proof if we'd used `Equivalence.mk`, -- but we choose to avoid using that here, for the sake of good structure projection `simp` -- lemmas. functor_unitIso_comp X := by dsimp - rw [← f.functor.map_comp_assoc, e.functor.map_comp, ← counitInv_app_functor, fun_inv_map, - Iso.inv_hom_id_app_assoc, assoc, Iso.inv_hom_id_app, counit_app_functor, ← Functor.map_comp] - erw [comp_id, Iso.hom_inv_id_app, Functor.map_id] + simp only [comp_id, id_comp, map_comp, fun_inv_map, comp_obj, id_obj, counitInv, + functor_unit_comp_assoc, assoc] + slice_lhs 2 3 => rw [← Functor.map_comp, Iso.inv_hom_id_app] + simp /-- Composing a functor with both functors of an equivalence yields a naturally isomorphic functor. -/ diff --git a/Mathlib/CategoryTheory/Localization/Predicate.lean b/Mathlib/CategoryTheory/Localization/Predicate.lean index bc2855fa278ac..4178f7ade448c 100644 --- a/Mathlib/CategoryTheory/Localization/Predicate.lean +++ b/Mathlib/CategoryTheory/Localization/Predicate.lean @@ -432,7 +432,9 @@ same `MorphismProperty C`, this is an equivalence of categories `D₁ ≌ D₂`. def uniq : D₁ ≌ D₂ := (equivalenceFromModel L₁ W').symm.trans (equivalenceFromModel L₂ W') -lemma uniq_symm : (uniq L₁ L₂ W').symm = uniq L₂ L₁ W' := rfl +lemma uniq_symm : (uniq L₁ L₂ W').symm = uniq L₂ L₁ W' := by + dsimp [uniq, Equivalence.trans] + ext <;> aesop /-- The functor of equivalence of localized categories given by `Localization.uniq` is compatible with the localization functors. -/ From a57f31753dde989bc415cb64277a1688866e443c Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Sun, 5 Jan 2025 17:28:34 +0000 Subject: [PATCH 167/189] chore(1000): remove nonexistent fields (#20493) `comment` is not a supported field and so the website build is [broken](https://github.com/leanprover-community/leanprover-community.github.io/actions/runs/12616366640/job/35157520815). Changes to the website to actually support this would be welcome. --- docs/1000.yaml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/docs/1000.yaml b/docs/1000.yaml index 64168a6a6f4ae..50d0e3fb4af8f 100644 --- a/docs/1000.yaml +++ b/docs/1000.yaml @@ -1329,7 +1329,7 @@ Q1191862: Q1194053: title: Metrization theorems decl: TopologicalSpace.metrizableSpace_of_t3_secondCountable - comment: "Urysohn's metrization theorem (only)" + # comment: "Urysohn's metrization theorem (only)" Q1196538: title: Descartes's theorem @@ -1384,7 +1384,7 @@ Q1306092: Q1306095: title: Whitney embedding theorem decl: exists_embedding_euclidean_of_compact - comment: "baby version: for compact manifolds; embedding into some *n*" + # comment: "baby version: for compact manifolds; embedding into some *n*" Q1307676: title: Künneth theorem From 95cdc7a6aba073b8cca8ba8d806e352b4e8ca0e5 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 5 Jan 2025 22:19:13 +0000 Subject: [PATCH 168/189] chore: replace `aesop` with `simp` where possible (#20483) I'm investigating how `aesop` is being used (and making comparisons with the in development `grind` tactic), but finding many uses of `aesop` can in fact just be replaced by `simp`. Hopefully using the simpler tool is both faster, and more readable. We'll see. --- .../Category/ModuleCat/Presheaf/Sheafify.lean | 2 +- .../Algebra/Category/ModuleCat/Subobject.lean | 3 +-- Mathlib/Algebra/Category/Semigrp/Basic.lean | 2 +- Mathlib/Algebra/Homology/Additive.lean | 2 +- Mathlib/Algebra/Homology/Augment.lean | 4 ++-- .../Homology/HomologicalBicomplex.lean | 4 ++-- .../Algebra/Homology/HomologySequence.lean | 4 ++-- .../HomotopyCategory/DegreewiseSplit.lean | 2 +- .../ShortComplex/FunctorEquivalence.lean | 8 +++---- .../ShortComplex/HomologicalComplex.lean | 2 +- .../Homology/ShortComplex/LeftHomology.lean | 2 +- .../Algebra/Homology/ShortComplex/Limits.lean | 4 ++-- .../ShortComplex/PreservesHomology.lean | 4 ++-- .../Homology/ShortComplex/RightHomology.lean | 6 ++--- .../Homology/ShortComplex/SnakeLemma.lean | 4 ++-- Mathlib/Algebra/Homology/Single.lean | 2 +- Mathlib/Algebra/Lie/Weights/Basic.lean | 2 +- .../Algebra/Module/Presentation/Basic.lean | 4 ++-- .../Algebra/Module/Submodule/Pointwise.lean | 4 ++-- .../PrimeSpectrum/Basic.lean | 2 +- Mathlib/AlgebraicTopology/CechNerve.lean | 2 +- .../SimplicialObject/Basic.lean | 2 +- .../ContinuousFunctionalCalculus/Unital.lean | 2 +- .../ContinuousFunctionalCalculus/ExpLog.lean | 2 +- .../ContinuousFunctionalCalculus/Rpow.lean | 2 +- Mathlib/Analysis/SpecificLimits/Basic.lean | 2 +- Mathlib/CategoryTheory/Abelian/Exact.lean | 4 ++-- .../Abelian/InjectiveResolution.lean | 2 +- Mathlib/CategoryTheory/Action/Limits.lean | 2 +- Mathlib/CategoryTheory/Action/Monoidal.lean | 2 +- Mathlib/CategoryTheory/Adjunction/Basic.lean | 4 ++-- .../Bicategory/FunctorBicategory/Oplax.lean | 6 ++--- Mathlib/CategoryTheory/Category/Cat.lean | 2 +- .../CategoryTheory/Category/Cat/Limit.lean | 4 ++-- Mathlib/CategoryTheory/Comma/Basic.lean | 4 ++-- .../CategoryTheory/Comma/Presheaf/Basic.lean | 6 ++--- .../Comma/StructuredArrow/Basic.lean | 18 +++++++------- Mathlib/CategoryTheory/DiscreteCategory.lean | 4 ++-- Mathlib/CategoryTheory/Elements.lean | 2 +- Mathlib/CategoryTheory/Equivalence.lean | 10 ++++---- Mathlib/CategoryTheory/Functor/Const.lean | 2 +- .../Functor/Derived/RightDerived.lean | 4 ++-- .../CategoryTheory/Functor/FunctorHom.lean | 2 +- .../Functor/KanExtension/Basic.lean | 4 ++-- .../Functor/KanExtension/Pointwise.lean | 8 +++---- .../CategoryTheory/Galois/Decomposition.lean | 2 +- Mathlib/CategoryTheory/GlueData.lean | 4 ++-- .../CategoryTheory/GradedObject/Single.lean | 2 +- Mathlib/CategoryTheory/Groupoid.lean | 2 +- .../Idempotents/FunctorExtension.lean | 4 ++-- .../Idempotents/HomologicalComplex.lean | 2 +- .../CategoryTheory/Idempotents/Karoubi.lean | 4 ++-- .../Idempotents/KaroubiKaroubi.lean | 2 +- Mathlib/CategoryTheory/IsConnected.lean | 2 +- Mathlib/CategoryTheory/Limits/Cones.lean | 6 ++--- .../Limits/Constructions/Filtered.lean | 4 ++-- .../Limits/Constructions/ZeroObjects.lean | 8 +++---- Mathlib/CategoryTheory/Limits/HasLimits.lean | 8 +++---- Mathlib/CategoryTheory/Limits/IsLimit.lean | 8 +++---- Mathlib/CategoryTheory/Limits/MonoCoprod.lean | 2 +- Mathlib/CategoryTheory/Limits/Opposites.lean | 14 +++++------ .../Limits/Preserves/Shapes/Biproducts.lean | 8 +++---- .../Limits/Shapes/Biproducts.lean | 10 ++++---- .../Limits/Shapes/Equalizers.lean | 20 ++++++++-------- .../CategoryTheory/Limits/Shapes/Images.lean | 2 +- .../Limits/Shapes/IsTerminal.lean | 24 +++++++++---------- .../CategoryTheory/Limits/Shapes/Kernels.lean | 10 ++++---- .../Limits/Shapes/Products.lean | 4 ++-- .../Limits/Shapes/Pullback/CommSq.lean | 16 ++++++------- .../Limits/Shapes/Pullback/Square.lean | 4 ++-- .../Limits/Shapes/Reflexive.lean | 2 +- .../Limits/Shapes/StrongEpi.lean | 4 ++-- .../Limits/Shapes/Terminal.lean | 8 +++---- .../CategoryTheory/Limits/Shapes/Types.lean | 2 +- .../Limits/Shapes/WidePullbacks.lean | 6 ++--- .../Localization/Bifunctor.lean | 4 ++-- .../Localization/Construction.lean | 2 +- .../DerivabilityStructure/Constructor.lean | 2 +- .../Monoidal/Braided/Basic.lean | 2 +- Mathlib/CategoryTheory/Monoidal/End.lean | 2 +- .../Monoidal/Internal/Limits.lean | 2 +- .../CategoryTheory/Monoidal/Subcategory.lean | 2 +- Mathlib/CategoryTheory/NatIso.lean | 2 +- Mathlib/CategoryTheory/Opposites.lean | 2 +- Mathlib/CategoryTheory/Pi/Basic.lean | 2 +- .../Preadditive/Biproducts.lean | 2 +- .../Preadditive/OfBiproducts.lean | 8 +++---- Mathlib/CategoryTheory/Simple.lean | 2 +- Mathlib/CategoryTheory/SingleObj.lean | 4 ++-- Mathlib/CategoryTheory/Sites/Pretopology.lean | 2 +- Mathlib/CategoryTheory/Sites/Subsheaf.lean | 2 +- Mathlib/CategoryTheory/Subobject/Comma.lean | 4 ++-- Mathlib/CategoryTheory/Subobject/Lattice.lean | 6 ++--- .../CategoryTheory/Triangulated/Basic.lean | 2 +- .../CategoryTheory/Triangulated/Functor.lean | 6 ++--- .../Opposite/Pretriangulated.lean | 6 ++--- .../Triangulated/Opposite/Triangle.lean | 6 ++--- .../Triangulated/Pretriangulated.lean | 6 ++--- .../Triangulated/Subcategory.lean | 4 ++-- .../Triangulated/TriangleShift.lean | 8 +++---- Mathlib/CategoryTheory/Yoneda.lean | 4 ++-- .../Computability/AkraBazzi/AkraBazzi.lean | 4 ++-- Mathlib/Data/MLList/BestFirst.lean | 2 +- Mathlib/Data/Matrix/ColumnRowPartitioned.lean | 3 +-- Mathlib/Data/Matroid/Closure.lean | 2 +- Mathlib/Data/PFunctor/Multivariate/M.lean | 2 +- Mathlib/Data/Part.lean | 2 +- .../LocallyRingedSpace/HasColimits.lean | 2 +- .../GroupTheory/SpecificGroups/KleinFour.lean | 2 +- Mathlib/LinearAlgebra/FreeProduct/Basic.lean | 2 +- .../LinearAlgebra/InvariantBasisNumber.lean | 4 ++-- Mathlib/LinearAlgebra/Multilinear/Basic.lean | 4 ++-- Mathlib/LinearAlgebra/Semisimple.lean | 3 +-- .../LinearAlgebra/TensorProduct/Quotient.lean | 4 ++-- Mathlib/ModelTheory/ElementaryMaps.lean | 10 ++++---- Mathlib/ModelTheory/Substructures.lean | 8 +++---- Mathlib/Order/BooleanSubalgebra.lean | 8 +++---- Mathlib/Probability/StrongLaw.lean | 2 +- .../GroupCohomology/LowDegree.lean | 8 +++---- .../AdicCompletion/Functoriality.lean | 2 +- .../Topology/Category/Profinite/AsLimit.lean | 2 +- Mathlib/Topology/ContinuousOn.lean | 2 +- .../Sheaves/SheafCondition/OpensLeCover.lean | 2 +- 123 files changed, 267 insertions(+), 270 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean index d1dec86a2c040..440f20249af9e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean @@ -318,7 +318,7 @@ noncomputable def sheafify : SheafOfModules.{v} R where def toSheafify : M₀ ⟶ (restrictScalars α).obj (sheafify α φ).val := homMk φ (fun X r₀ m₀ ↦ by simpa using (Sheafify.map_smul_eq α φ (α.app _ r₀) (φ.app _ m₀) (𝟙 _) - r₀ (by aesop) m₀ (by simp)).symm) + r₀ (by simp) m₀ (by simp)).symm) lemma toSheafify_app_apply (X : Cᵒᵖ) (x : M₀.obj X) : ((toSheafify α φ).app X).hom x = φ.app X x := rfl diff --git a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean index 8b5e532a67127..5fb8394a9c465 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean @@ -85,8 +85,7 @@ theorem toKernelSubobject_arrow {M N : ModuleCat R} {f : M ⟶ N} (x : LinearMap suffices ((arrow ((kernelSubobject f))) ∘ (kernelSubobjectIso f ≪≫ kernelIsoKer f).inv) x = x by convert this rw [Iso.trans_inv, ← LinearMap.coe_comp, ← hom_comp, Category.assoc] - simp only [Category.assoc, kernelSubobject_arrow', kernelIsoKer_inv_kernel_ι] - aesop_cat + simp /-- An extensionality lemma showing that two elements of a cokernel by an image are equal if they differ by an element of the image. diff --git a/Mathlib/Algebra/Category/Semigrp/Basic.lean b/Mathlib/Algebra/Category/Semigrp/Basic.lean index d30f2b19f232e..e6e54e5ccab88 100644 --- a/Mathlib/Algebra/Category/Semigrp/Basic.lean +++ b/Mathlib/Algebra/Category/Semigrp/Basic.lean @@ -44,7 +44,7 @@ namespace MagmaCat @[to_additive] instance bundledHom : BundledHom @MulHom := ⟨@MulHom.toFun, @MulHom.id, @MulHom.comp, - by intros; apply @DFunLike.coe_injective, by aesop_cat, by aesop_cat⟩ + by intros; apply @DFunLike.coe_injective, by aesop_cat, by simp⟩ -- Porting note: deriving failed for `ConcreteCategory`, -- "default handlers have not been implemented yet" diff --git a/Mathlib/Algebra/Homology/Additive.lean b/Mathlib/Algebra/Homology/Additive.lean index 1b14c45bf5b00..86d5e417b4051 100644 --- a/Mathlib/Algebra/Homology/Additive.lean +++ b/Mathlib/Algebra/Homology/Additive.lean @@ -176,7 +176,7 @@ theorem NatTrans.mapHomologicalComplex_naturality {c : ComplexShape ι} {F G : W (α : F ⟶ G) {C D : HomologicalComplex W₁ c} (f : C ⟶ D) : (F.mapHomologicalComplex c).map f ≫ (NatTrans.mapHomologicalComplex α c).app D = (NatTrans.mapHomologicalComplex α c).app C ≫ (G.mapHomologicalComplex c).map f := by - aesop_cat + simp /-- A natural isomorphism between functors induces a natural isomorphism between those functors applied to homological complexes. diff --git a/Mathlib/Algebra/Homology/Augment.lean b/Mathlib/Algebra/Homology/Augment.lean index b7701f75f0d90..2f19bc16f99bb 100644 --- a/Mathlib/Algebra/Homology/Augment.lean +++ b/Mathlib/Algebra/Homology/Augment.lean @@ -37,7 +37,7 @@ The components of this chain map are `C.d 1 0` in degree 0, and zero otherwise. -/ def truncateTo [HasZeroObject V] [HasZeroMorphisms V] (C : ChainComplex V ℕ) : truncate.obj C ⟶ (single₀ V).obj (C.X 0) := - (toSingle₀Equiv (truncate.obj C) (C.X 0)).symm ⟨C.d 1 0, by aesop⟩ + (toSingle₀Equiv (truncate.obj C) (C.X 0)).symm ⟨C.d 1 0, by simp⟩ -- PROJECT when `V` is abelian (but not generally?) -- `[∀ n, Exact (C.d (n+2) (n+1)) (C.d (n+1) n)] [Epi (C.d 1 0)]` iff `QuasiIso (C.truncate_to)` @@ -204,7 +204,7 @@ The components of this chain map are `C.d 0 1` in degree 0, and zero otherwise. -/ def toTruncate [HasZeroObject V] [HasZeroMorphisms V] (C : CochainComplex V ℕ) : (single₀ V).obj (C.X 0) ⟶ truncate.obj C := - (fromSingle₀Equiv (truncate.obj C) (C.X 0)).symm ⟨C.d 0 1, by aesop⟩ + (fromSingle₀Equiv (truncate.obj C) (C.X 0)).symm ⟨C.d 0 1, by simp⟩ variable [HasZeroMorphisms V] diff --git a/Mathlib/Algebra/Homology/HomologicalBicomplex.lean b/Mathlib/Algebra/Homology/HomologicalBicomplex.lean index 634f2df3056f2..e9f7ce59fdfb9 100644 --- a/Mathlib/Algebra/Homology/HomologicalBicomplex.lean +++ b/Mathlib/Algebra/Homology/HomologicalBicomplex.lean @@ -177,7 +177,7 @@ def flipEquivalenceUnitIso : 𝟭 (HomologicalComplex₂ C c₁ c₂) ≅ flipFunctor C c₁ c₂ ⋙ flipFunctor C c₂ c₁ := NatIso.ofComponents (fun K => HomologicalComplex.Hom.isoOfComponents (fun i₁ => HomologicalComplex.Hom.isoOfComponents (fun _ => Iso.refl _) - (by aesop_cat)) (by aesop_cat)) (by aesop_cat) + (by simp)) (by aesop_cat)) (by aesop_cat) /-- Auxiliary definition for `HomologicalComplex₂.flipEquivalence`. -/ @[simps!] @@ -185,7 +185,7 @@ def flipEquivalenceCounitIso : flipFunctor C c₂ c₁ ⋙ flipFunctor C c₁ c₂ ≅ 𝟭 (HomologicalComplex₂ C c₂ c₁) := NatIso.ofComponents (fun K => HomologicalComplex.Hom.isoOfComponents (fun i₂ => HomologicalComplex.Hom.isoOfComponents (fun _ => Iso.refl _) - (by aesop_cat)) (by aesop_cat)) (by aesop_cat) + (by simp)) (by aesop_cat)) (by aesop_cat) /-- Flipping a complex of complexes over the diagonal, as an equivalence of categories. -/ @[simps] diff --git a/Mathlib/Algebra/Homology/HomologySequence.lean b/Mathlib/Algebra/Homology/HomologySequence.lean index 6a325db2f53a9..c3743d891f42f 100644 --- a/Mathlib/Algebra/Homology/HomologySequence.lean +++ b/Mathlib/Algebra/Homology/HomologySequence.lean @@ -165,7 +165,7 @@ noncomputable def composableArrows₃Functor [CategoryWithHomology C] : HomologicalComplex C c ⥤ ComposableArrows C 3 where obj K := composableArrows₃ K i j map {K L} φ := ComposableArrows.homMk₃ (homologyMap φ i) (opcyclesMap φ i) (cyclesMap φ j) - (homologyMap φ j) (by aesop_cat) (by aesop_cat) (by aesop_cat) + (homologyMap φ j) (by simp) (by simp) (by simp) end HomologySequence @@ -311,7 +311,7 @@ lemma homology_exact₂ : (ShortComplex.mk (HomologicalComplex.homologyMap S.f i have e : S.map (HomologicalComplex.homologyFunctor C c i) ≅ S.map (HomologicalComplex.opcyclesFunctor C c i) := ShortComplex.isoMk (asIso (S.X₁.homologyι i)) - (asIso (S.X₂.homologyι i)) (asIso (S.X₃.homologyι i)) (by aesop_cat) (by aesop_cat) + (asIso (S.X₂.homologyι i)) (asIso (S.X₃.homologyι i)) (by simp) (by simp) exact ShortComplex.exact_of_iso e.symm (opcycles_right_exact S hS.exact i) /-- Exactness of `S.X₂.homology i ⟶ S.X₃.homology i ⟶ S.X₁.homology j`. -/ diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean index cabff6665b54e..c454aa0ad7d69 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean @@ -201,7 +201,7 @@ noncomputable def triangleRotateIsoTriangleOfDegreewiseSplit : (triangle φ).rotate ≅ triangleOfDegreewiseSplit _ (triangleRotateShortComplexSplitting φ) := Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) - (by aesop_cat) (by aesop_cat) (by ext; dsimp; simp) + (by simp) (by simp) (by ext; dsimp; simp) /-- The triangle `(triangleh φ).rotate` is isomorphic to a triangle attached to a degreewise split short exact sequence of cochain complexes. -/ diff --git a/Mathlib/Algebra/Homology/ShortComplex/FunctorEquivalence.lean b/Mathlib/Algebra/Homology/ShortComplex/FunctorEquivalence.lean index 3e5b25181998c..b16976a6c8d8f 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/FunctorEquivalence.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/FunctorEquivalence.lean @@ -50,9 +50,9 @@ def inverse : (J ⥤ ShortComplex C) ⥤ ShortComplex (J ⥤ C) where @[simps!] def unitIso : 𝟭 _ ≅ functor J C ⋙ inverse J C := NatIso.ofComponents (fun _ => isoMk - (NatIso.ofComponents (fun _ => Iso.refl _) (by aesop_cat)) - (NatIso.ofComponents (fun _ => Iso.refl _) (by aesop_cat)) - (NatIso.ofComponents (fun _ => Iso.refl _) (by aesop_cat)) + (NatIso.ofComponents (fun _ => Iso.refl _) (by simp)) + (NatIso.ofComponents (fun _ => Iso.refl _) (by simp)) + (NatIso.ofComponents (fun _ => Iso.refl _) (by simp)) (by aesop_cat) (by aesop_cat)) (by aesop_cat) /-- The counit isomorphism of the equivalence @@ -61,7 +61,7 @@ def unitIso : 𝟭 _ ≅ functor J C ⋙ inverse J C := def counitIso : inverse J C ⋙ functor J C ≅ 𝟭 _ := NatIso.ofComponents (fun _ => NatIso.ofComponents (fun _ => isoMk (Iso.refl _) (Iso.refl _) (Iso.refl _) - (by aesop_cat) (by aesop_cat)) (by aesop_cat)) (by aesop_cat) + (by simp) (by simp)) (by aesop_cat)) (by aesop_cat) end FunctorEquivalence diff --git a/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean b/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean index 415cac78ef15f..97bd4dc2b02fa 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean @@ -50,7 +50,7 @@ when `c.prev j = i` and `c.next j = k`. -/ noncomputable def natIsoSc' (i j k : ι) (hi : c.prev j = i) (hk : c.next j = k) : shortComplexFunctor C c j ≅ shortComplexFunctor' C c i j k := NatIso.ofComponents (fun K => ShortComplex.isoMk (K.XIsoOfEq hi) (Iso.refl _) (K.XIsoOfEq hk) - (by aesop_cat) (by aesop_cat)) (by aesop_cat) + (by simp) (by simp)) (by aesop_cat) variable {C c} diff --git a/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean b/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean index a8abcedd4972c..8778624055502 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean @@ -1008,7 +1008,7 @@ lemma hasCokernel [S.HasLeftHomology] [HasKernel S.g] : haveI : HasColimit (parallelPair h.f' 0) := ⟨⟨⟨_, h.hπ'⟩⟩⟩ let e : parallelPair (kernel.lift S.g S.f S.zero) 0 ≅ parallelPair h.f' 0 := parallelPair.ext (Iso.refl _) (IsLimit.conePointUniqueUpToIso (kernelIsKernel S.g) h.hi) - (by aesop_cat) (by aesop_cat) + (by aesop_cat) (by simp) exact hasColimitOfIso e end HasLeftHomology diff --git a/Mathlib/Algebra/Homology/ShortComplex/Limits.lean b/Mathlib/Algebra/Homology/ShortComplex/Limits.lean index 1f94e0fc71c5e..a7a46d412c3f3 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Limits.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Limits.lean @@ -66,7 +66,7 @@ noncomputable def limitCone : Cone F := Cone.mk (ShortComplex.mk (limMap (whiskerLeft F π₁Toπ₂)) (limMap (whiskerLeft F π₂Toπ₃)) (by aesop_cat)) { app := fun j => Hom.mk (limit.π _ _) (limit.π _ _) (limit.π _ _) - (by aesop_cat) (by aesop_cat) + (by simp) (by simp) naturality := fun _ _ f => by ext all_goals @@ -199,7 +199,7 @@ noncomputable def colimitCocone : Cocone F := Cocone.mk (ShortComplex.mk (colimMap (whiskerLeft F π₁Toπ₂)) (colimMap (whiskerLeft F π₂Toπ₃)) (by aesop_cat)) { app := fun j => Hom.mk (colimit.ι (F ⋙ π₁) _) (colimit.ι (F ⋙ π₂) _) - (colimit.ι (F ⋙ π₃) _) (by aesop_cat) (by aesop_cat) + (colimit.ι (F ⋙ π₃) _) (by simp) (by simp) naturality := fun _ _ f => by ext · dsimp; erw [comp_id, colimit.w (F ⋙ π₁)] diff --git a/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean b/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean index eacf12a85c953..df003423f5f9d 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean @@ -834,7 +834,7 @@ lemma preservesLeftHomology_of_zero_g (hg : S.g = 0) f' := by have := h.isIso_i hg let e : parallelPair h.f' 0 ≅ parallelPair S.f 0 := - parallelPair.ext (Iso.refl _) (asIso h.i) (by aesop_cat) (by aesop_cat) + parallelPair.ext (Iso.refl _) (asIso h.i) (by simp) (by simp) exact Limits.preservesColimit_of_iso_diagram F e.symm}⟩ /-- If a short complex `S` is such that `S.f = 0` and that the kernel of `S.g` is preserved @@ -848,7 +848,7 @@ lemma preservesRightHomology_of_zero_f (hf : S.f = 0) g' := by have := h.isIso_p hf let e : parallelPair S.g 0 ≅ parallelPair h.g' 0 := - parallelPair.ext (asIso h.p) (Iso.refl _) (by aesop_cat) (by aesop_cat) + parallelPair.ext (asIso h.p) (Iso.refl _) (by simp) (by simp) exact Limits.preservesLimit_of_iso_diagram F e }⟩ end Functor diff --git a/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean b/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean index 9832970621a74..b6c4767c9e7b7 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean @@ -150,7 +150,7 @@ def ofIsLimitKernelFork (hf : S.f = 0) (c : KernelFork S.g) (hc : IsLimit c) : wp := by rw [comp_id, hf] hp := CokernelCofork.IsColimit.ofId _ hf wι := KernelFork.condition _ - hι := IsLimit.ofIsoLimit hc (Fork.ext (Iso.refl _) (by aesop_cat)) + hι := IsLimit.ofIsoLimit hc (Fork.ext (Iso.refl _) (by simp)) @[simp] lemma ofIsLimitKernelFork_g' (hf : S.f = 0) (c : KernelFork S.g) (hc : IsLimit c) : (ofIsLimitKernelFork S hf c hc).g' = S.g := by @@ -173,7 +173,7 @@ def ofIsColimitCokernelCofork (hg : S.g = 0) (c : CokernelCofork S.f) (hc : IsCo p := c.π ι := 𝟙 _ wp := CokernelCofork.condition _ - hp := IsColimit.ofIsoColimit hc (Cofork.ext (Iso.refl _) (by aesop_cat)) + hp := IsColimit.ofIsoColimit hc (Cofork.ext (Iso.refl _) (by simp)) wι := Cofork.IsColimit.hom_ext hc (by simp [hg]) hι := KernelFork.IsLimit.ofId _ (Cofork.IsColimit.hom_ext hc (by simp [hg])) @@ -1287,7 +1287,7 @@ lemma hasKernel [S.HasRightHomology] [HasCokernel S.f] : haveI : HasLimit (parallelPair h.g' 0) := ⟨⟨⟨_, h.hι'⟩⟩⟩ let e : parallelPair (cokernel.desc S.f S.g S.zero) 0 ≅ parallelPair h.g' 0 := parallelPair.ext (IsColimit.coconePointUniqueUpToIso (colimit.isColimit _) h.hp) - (Iso.refl _) (coequalizer.hom_ext (by simp)) (by aesop_cat) + (Iso.refl _) (coequalizer.hom_ext (by simp)) (by simp) exact hasLimitOfIso e.symm end HasRightHomology diff --git a/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean b/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean index 8fe3dbac06932..c3abf86cbb643 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean @@ -345,7 +345,7 @@ lemma op_δ : S.op.δ = S.δ.op := Quiver.Hom.unop_inj (by /-- The duality isomorphism `S.L₂'.op ≅ S.op.L₁'`. -/ noncomputable def L₂'OpIso : S.L₂'.op ≅ S.op.L₁' := - ShortComplex.isoMk (Iso.refl _) (Iso.refl _) (Iso.refl _) (by aesop_cat) + ShortComplex.isoMk (Iso.refl _) (Iso.refl _) (Iso.refl _) (by simp) (by dsimp; simp only [id_comp, comp_id, S.op_δ]) /-- Exactness of `L₀.X₃ ⟶ L₃.X₁ ⟶ L₃.X₂`. -/ @@ -475,7 +475,7 @@ noncomputable def functorP : SnakeInput C ⥤ C where obj S := S.P map f := pullback.map _ _ _ _ f.f₁.τ₂ f.f₀.τ₃ f.f₁.τ₃ f.f₁.comm₂₃.symm (congr_arg ShortComplex.Hom.τ₃ f.comm₀₁.symm) - map_id _ := by dsimp [P]; aesop_cat + map_id _ := by dsimp [P]; simp map_comp _ _ := by dsimp [P]; aesop_cat @[reassoc] diff --git a/Mathlib/Algebra/Homology/Single.lean b/Mathlib/Algebra/Homology/Single.lean index a9cbef23cd3c9..b6dc763c87a95 100644 --- a/Mathlib/Algebra/Homology/Single.lean +++ b/Mathlib/Algebra/Homology/Single.lean @@ -220,7 +220,7 @@ noncomputable def toSingle₀Equiv (C : ChainComplex V ℕ) (X : V) : obtain rfl : i = 1 := by simpa using hi.symm exact f.2) left_inv φ := by aesop_cat - right_inv f := by aesop_cat + right_inv f := by simp @[simp] lemma toSingle₀Equiv_symm_apply_f_zero {C : ChainComplex V ℕ} {X : V} diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index 7ecb88285cfc7..9cb657c719f2d 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -228,7 +228,7 @@ variable {M} @[ext] lemma ext {χ₁ χ₂ : Weight R L M} (h : ∀ x, χ₁ x = χ₂ x) : χ₁ = χ₂ := by cases' χ₁ with f₁ _; cases' χ₂ with f₂ _; aesop -lemma ext_iff' {χ₁ χ₂ : Weight R L M} : (χ₁ : L → R) = χ₂ ↔ χ₁ = χ₂ := by aesop +lemma ext_iff' {χ₁ χ₂ : Weight R L M} : (χ₁ : L → R) = χ₂ ↔ χ₁ = χ₂ := by simp lemma exists_ne_zero (χ : Weight R L M) : ∃ x ∈ genWeightSpace M χ, x ≠ 0 := by diff --git a/Mathlib/Algebra/Module/Presentation/Basic.lean b/Mathlib/Algebra/Module/Presentation/Basic.lean index b12188e759ad7..e9c9735f6af10 100644 --- a/Mathlib/Algebra/Module/Presentation/Basic.lean +++ b/Mathlib/Algebra/Module/Presentation/Basic.lean @@ -354,8 +354,8 @@ certain linear equations. -/ noncomputable def linearMapEquiv : (M →ₗ[A] N) ≃ relations.Solution N where toFun f := solution.postcomp f invFun s := h.desc s - left_inv f := h.postcomp_injective (by aesop) - right_inv s := by aesop + left_inv f := h.postcomp_injective (by simp) + right_inv s := by simp section diff --git a/Mathlib/Algebra/Module/Submodule/Pointwise.lean b/Mathlib/Algebra/Module/Submodule/Pointwise.lean index bad57cad693c0..b3c2fa905b8d3 100644 --- a/Mathlib/Algebra/Module/Submodule/Pointwise.lean +++ b/Mathlib/Algebra/Module/Submodule/Pointwise.lean @@ -475,8 +475,8 @@ lemma mem_singleton_set_smul [SMulCommClass R S M] (r : S) (x : M) : exact ⟨t • n, by aesop, smul_comm _ _ _⟩ · rcases h₁ with ⟨m₁, h₁, rfl⟩ rcases h₂ with ⟨m₂, h₂, rfl⟩ - exact ⟨m₁ + m₂, Submodule.add_mem _ h₁ h₂, by aesop⟩ - · exact ⟨0, Submodule.zero_mem _, by aesop⟩ + exact ⟨m₁ + m₂, Submodule.add_mem _ h₁ h₂, by simp⟩ + · exact ⟨0, Submodule.zero_mem _, by simp⟩ · aesop lemma smul_inductionOn_pointwise [SMulCommClass S R M] {a : S} {p : (x : M) → x ∈ a • N → Prop} diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 680ed47f8dce6..2eabfbfe9f773 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -142,7 +142,7 @@ theorem t1Space_iff_isField [IsDomain R] : T1Space (PrimeSpectrum R) ↔ IsField (mt (Ring.ne_bot_of_isMaximal_of_not_isField <| (isClosed_singleton_iff_isMaximal _).1 (T1Space.t1 ⟨⊥, hbot⟩)) - (by aesop)) + (by simp)) · refine ⟨fun x => (isClosed_singleton_iff_isMaximal x).2 ?_⟩ by_cases hx : x.asIdeal = ⊥ · letI := h.toSemifield diff --git a/Mathlib/AlgebraicTopology/CechNerve.lean b/Mathlib/AlgebraicTopology/CechNerve.lean index 667421531fc21..1a6755cccda6b 100644 --- a/Mathlib/AlgebraicTopology/CechNerve.lean +++ b/Mathlib/AlgebraicTopology/CechNerve.lean @@ -50,7 +50,7 @@ variable [∀ n : ℕ, HasWidePullback.{0} f.right (fun _ : Fin (n + 1) => f.lef def cechNerve : SimplicialObject C where obj n := widePullback.{0} f.right (fun _ : Fin (n.unop.len + 1) => f.left) fun _ => f.hom map g := WidePullback.lift (WidePullback.base _) - (fun i => WidePullback.π _ (g.unop.toOrderHom i)) (by aesop_cat) + (fun i => WidePullback.π _ (g.unop.toOrderHom i)) (by simp) /-- The morphism between Čech nerves associated to a morphism of arrows. -/ @[simps] diff --git a/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean index d8a661f4056b0..ae05211a3fe1e 100644 --- a/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean @@ -834,7 +834,7 @@ object and back is isomorphic to the given object. -/ @[simps!] def SimplicialObject.Augmented.rightOpLeftOpIso (X : SimplicialObject.Augmented C) : X.rightOp.leftOp ≅ X := - Comma.isoMk X.left.rightOpLeftOpIso (CategoryTheory.eqToIso <| by aesop_cat) + Comma.isoMk X.left.rightOpLeftOpIso (CategoryTheory.eqToIso <| by simp) /-- Converting an augmented cosimplicial object to an augmented simplicial object and back is isomorphic to the given object. -/ diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index ea55c06fddf1e..0cf2fabda7a17 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -724,7 +724,7 @@ variable [Ring A] [StarRing A] [Algebra R A] [ContinuousFunctionalCalculus R p] 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 + simp alias ⟨_, isUnit_cfc⟩ := isUnit_cfc_iff diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean index 27a9142d793a9..bdb5875943ebe 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean @@ -152,7 +152,7 @@ lemma log_pow (n : ℕ) (a : A) (ha₂ : ∀ x ∈ spectrum ℝ a, 0 < x) 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) + have hcont : ContinuousOn Real.log (Real.exp '' spectrum ℝ a) := by fun_prop (disch := simp) rw [log, ← real_exp_eq_normedSpace_exp, ← cfc_comp' Real.log Real.exp a hcont] simp [cfc_id' (R := ℝ) a] diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean index 487d14cd2871b..0bdd483e3237a 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean @@ -365,7 +365,7 @@ lemma sqrt_rpow_nnreal {a : A} {x : ℝ≥0} : sqrt (a ^ (x : ℝ)) = a ^ (x / 2 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₁ : 0 < x := lt_of_le_of_ne (by simp) (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)] diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index c226f8f39d7af..40a14b4c0240c 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -109,7 +109,7 @@ theorem Filter.EventuallyEq.div_mul_cancel_atTop {α K : Type*} [LinearOrderedSe {f g : α → K} {l : Filter α} (hg : Tendsto g l atTop) : (fun x ↦ f x / g x * g x) =ᶠ[l] fun x ↦ f x := div_mul_cancel <| hg.mono_right <| le_principal_iff.mpr <| - mem_of_superset (Ioi_mem_atTop 0) <| by aesop + mem_of_superset (Ioi_mem_atTop 0) <| by simp /-- If when `x` tends to `∞`, `g` tends to `∞` and `f x / g x` tends to a positive constant, then `f` tends to `∞`. -/ diff --git a/Mathlib/CategoryTheory/Abelian/Exact.lean b/Mathlib/CategoryTheory/Abelian/Exact.lean index badc3ad7a0d10..30b6b7aab74df 100644 --- a/Mathlib/CategoryTheory/Abelian/Exact.lean +++ b/Mathlib/CategoryTheory/Abelian/Exact.lean @@ -105,7 +105,7 @@ def Exact.isLimitImage (h : S.Exact) : rw [exact_iff_kernel_ι_comp_cokernel_π_zero] at h exact KernelFork.IsLimit.ofι _ _ (fun u hu ↦ kernel.lift (cokernel.π S.f) u - (by rw [← kernel.lift_ι S.g u hu, Category.assoc, h, comp_zero])) (by aesop_cat) + (by rw [← kernel.lift_ι S.g u hu, Category.assoc, h, comp_zero])) (by simp) (fun _ _ _ hm => by rw [← cancel_mono (Abelian.image.ι S.f), hm, kernel.lift_ι]) /-- If `(f, g)` is exact, then `image.ι f` is a kernel of `g`. -/ @@ -123,7 +123,7 @@ def Exact.isColimitCoimage (h : S.Exact) : refine CokernelCofork.IsColimit.ofπ _ _ (fun u hu => cokernel.desc (kernel.ι S.g) u (by rw [← cokernel.π_desc S.f u hu, ← Category.assoc, h, zero_comp])) - (by aesop_cat) ?_ + (by simp) ?_ intros _ _ _ _ hm ext rw [hm, cokernel.π_desc] diff --git a/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean b/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean index 1473b0ef27413..9110373e9c67a 100644 --- a/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean +++ b/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean @@ -240,7 +240,7 @@ lemma InjectiveResolution.iso_hom_naturality {X Y : C} (f : X ⟶ Y) I.iso.hom ≫ (HomotopyCategory.quotient _ _).map φ := by apply HomotopyCategory.eq_of_homotopy apply descHomotopy f - all_goals aesop_cat + all_goals aesop @[reassoc] lemma InjectiveResolution.iso_inv_naturality {X Y : C} (f : X ⟶ Y) diff --git a/Mathlib/CategoryTheory/Action/Limits.lean b/Mathlib/CategoryTheory/Action/Limits.lean index 00fbf901e87fc..a997429802737 100644 --- a/Mathlib/CategoryTheory/Action/Limits.lean +++ b/Mathlib/CategoryTheory/Action/Limits.lean @@ -209,7 +209,7 @@ variable [HasZeroMorphisms V] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/10688): in order to ease automation, the `Zero` instance is introduced separately, -- and the lemma `Action.zero_hom` was moved just below -instance {X Y : Action V G} : Zero (X ⟶ Y) := ⟨0, by aesop_cat⟩ +instance {X Y : Action V G} : Zero (X ⟶ Y) := ⟨0, by simp⟩ @[simp] theorem zero_hom {X Y : Action V G} : (0 : X ⟶ Y).hom = 0 := diff --git a/Mathlib/CategoryTheory/Action/Monoidal.lean b/Mathlib/CategoryTheory/Action/Monoidal.lean index 5428444039ef2..98d9b5620c1a2 100644 --- a/Mathlib/CategoryTheory/Action/Monoidal.lean +++ b/Mathlib/CategoryTheory/Action/Monoidal.lean @@ -89,7 +89,7 @@ variable [BraidedCategory V] instance : BraidedCategory (Action V G) := braidedCategoryOfFaithful (Action.forget V G) (fun X Y => mkIso (β_ _ _) - (fun g => by simp [FunctorCategoryEquivalence.inverse])) (by aesop_cat) + (fun g => by simp [FunctorCategoryEquivalence.inverse])) (by simp) /-- When `V` is braided the forgetful functor `Action V G` to `V` is braided. -/ instance : (Action.forget V G).Braided where diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index 49a73af0c7d52..6c4183721f9f5 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -293,14 +293,14 @@ theorem eq_homEquiv_apply {A : C} {B : D} (f : F.obj A ⟶ B) (g : A ⟶ G.obj B def corepresentableBy (X : C) : (G ⋙ coyoneda.obj (Opposite.op X)).CorepresentableBy (F.obj X) where homEquiv := adj.homEquiv _ _ - homEquiv_comp := by aesop_cat + homEquiv_comp := by simp /-- If `adj : F ⊣ G`, and `Y : D`, then `G.obj Y` represents `X ↦ (F.obj X ⟶ Y)`-/ @[simps] def representableBy (Y : D) : (F.op ⋙ yoneda.obj Y).RepresentableBy (G.obj Y) where homEquiv := (adj.homEquiv _ _).symm - homEquiv_comp := by aesop_cat + homEquiv_comp := by simp end diff --git a/Mathlib/CategoryTheory/Bicategory/FunctorBicategory/Oplax.lean b/Mathlib/CategoryTheory/Bicategory/FunctorBicategory/Oplax.lean index 02482d12bfd09..a0058e9be01b0 100644 --- a/Mathlib/CategoryTheory/Bicategory/FunctorBicategory/Oplax.lean +++ b/Mathlib/CategoryTheory/Bicategory/FunctorBicategory/Oplax.lean @@ -50,19 +50,19 @@ def whiskerRight {η θ : F ⟶ G} (Γ : η ⟶ θ) (ι : G ⟶ H) : η ≫ ι -- Porting note: verified that projections are correct and changed @[simps] to @[simps!] @[simps!] def associator (η : F ⟶ G) (θ : G ⟶ H) (ι : H ⟶ I) : (η ≫ θ) ≫ ι ≅ η ≫ θ ≫ ι := - ModificationIso.ofComponents (fun a => α_ (η.app a) (θ.app a) (ι.app a)) (by aesop_cat) + ModificationIso.ofComponents (fun a => α_ (η.app a) (θ.app a) (ι.app a)) (by simp) /-- Left unitor for the vertical composition of oplax natural transformations. -/ -- Porting note: verified that projections are correct and changed @[simps] to @[simps!] @[simps!] def leftUnitor (η : F ⟶ G) : 𝟙 F ≫ η ≅ η := - ModificationIso.ofComponents (fun a => λ_ (η.app a)) (by aesop_cat) + ModificationIso.ofComponents (fun a => λ_ (η.app a)) (by simp) /-- Right unitor for the vertical composition of oplax natural transformations. -/ -- Porting note: verified that projections are correct and changed @[simps] to @[simps!] @[simps!] def rightUnitor (η : F ⟶ G) : η ≫ 𝟙 G ≅ η := - ModificationIso.ofComponents (fun a => ρ_ (η.app a)) (by aesop_cat) + ModificationIso.ofComponents (fun a => ρ_ (η.app a)) (by simp) end OplaxNatTrans diff --git a/Mathlib/CategoryTheory/Category/Cat.lean b/Mathlib/CategoryTheory/Category/Cat.lean index 7d16c775e8b1d..b76adff0b2f61 100644 --- a/Mathlib/CategoryTheory/Category/Cat.lean +++ b/Mathlib/CategoryTheory/Category/Cat.lean @@ -182,7 +182,7 @@ def typeToCat : Type u ⥤ Cat where simp only [id_eq, eqToHom_refl, Cat.id_map, Category.comp_id, Category.id_comp] apply ULift.ext aesop_cat - · aesop_cat + · simp map_comp f g := by apply Functor.ext; aesop_cat instance : Functor.Faithful typeToCat.{u} where diff --git a/Mathlib/CategoryTheory/Category/Cat/Limit.lean b/Mathlib/CategoryTheory/Category/Cat/Limit.lean index e16a6957e1e3f..ae7ce8e214d0b 100644 --- a/Mathlib/CategoryTheory/Category/Cat/Limit.lean +++ b/Mathlib/CategoryTheory/Category/Cat/Limit.lean @@ -67,10 +67,10 @@ instance (F : J ⥤ Cat.{v, v}) : Category (limit (F ⋙ Cat.objects)) where ← congr_fun (limit.w (homDiagram Y Z) h) g] id_comp _ := by apply Types.limit_ext.{v, v} - aesop_cat + simp comp_id _ := by apply Types.limit_ext.{v, v} - aesop_cat + simp /-- Auxiliary definition: the limit category. -/ @[simps] diff --git a/Mathlib/CategoryTheory/Comma/Basic.lean b/Mathlib/CategoryTheory/Comma/Basic.lean index ccbfebe05bd1b..ee0bc05acbe47 100644 --- a/Mathlib/CategoryTheory/Comma/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/Basic.lean @@ -289,7 +289,7 @@ theorem map_fst : map α β ⋙ fst L' R' = 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) + NatIso.ofComponents (fun _ => Iso.refl _) (by simp) /-- The equality between `map α β ⋙ snd L' R'` and `snd L R ⋙ F₂`, where `β : R ⋙ F ⟶ F₂ ⋙ R'`. -/ @@ -301,7 +301,7 @@ theorem map_snd : map α β ⋙ snd L' R' = 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) + NatIso.ofComponents (fun _ => Iso.refl _) (by simp) end diff --git a/Mathlib/CategoryTheory/Comma/Presheaf/Basic.lean b/Mathlib/CategoryTheory/Comma/Presheaf/Basic.lean index 56df6eace290c..94aa947d47cc9 100644 --- a/Mathlib/CategoryTheory/Comma/Presheaf/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/Presheaf/Basic.lean @@ -190,7 +190,7 @@ lemma yonedaArrow_val {Y : C} {η : yoneda.obj Y ⟶ A} {X : C} {s : yoneda.obj /-- If `η` is also `yoneda`-costructured, then `OverArrows η s` is just morphisms of costructured arrows. -/ def costructuredArrowIso (s t : CostructuredArrow yoneda A) : OverArrows s.hom t.hom ≅ t ⟶ s where - hom p := CostructuredArrow.homMk p.val (by aesop_cat) + hom p := CostructuredArrow.homMk p.val (by simp) inv f := yonedaArrow f.left f.w end OverArrows @@ -436,7 +436,7 @@ lemma app_unitForward {F : Cᵒᵖ ⥤ Type v} (η : F ⟶ A) (X : Cᵒᵖ) /-- Backward direction of the unit. -/ def unitBackward {F : Cᵒᵖ ⥤ Type v} (η : F ⟶ A) (X : C) : F.obj (op X) → YonedaCollection (restrictedYonedaObj η) X := - fun x => YonedaCollection.mk (yonedaEquiv.symm (η.app _ x)) ⟨x, ⟨by aesop_cat⟩⟩ + fun x => YonedaCollection.mk (yonedaEquiv.symm (η.app _ x)) ⟨x, ⟨by simp⟩⟩ lemma unitForward_unitBackward {F : Cᵒᵖ ⥤ Type v} (η : F ⟶ A) (X : C) : unitForward η X ∘ unitBackward η X = id := @@ -513,7 +513,7 @@ lemma counitForward_naturality₂ (s t : (CostructuredArrow yoneda A)ᵒᵖ) (f have : (CostructuredArrow.mkPrecomp t.unop.hom f.unop.left).op = f ≫ eqToHom (by simp [← CostructuredArrow.eq_mk]) := by apply Quiver.Hom.unop_inj - aesop_cat + simp aesop_cat /-- Backward direction of the counit. -/ diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean index 22de055a81308..42ce3a6553dbd 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean @@ -126,7 +126,7 @@ lemma homMk'_id (f : StructuredArrow S T) : homMk' f (𝟙 f.right) = eqToHom (b ext simp [eqToHom_right] -lemma homMk'_mk_id (f : S ⟶ T.obj Y) : homMk' (mk f) (𝟙 Y) = eqToHom (by aesop_cat) := +lemma homMk'_mk_id (f : S ⟶ T.obj Y) : homMk' (mk f) (𝟙 Y) = eqToHom (by simp) := homMk'_id _ lemma homMk'_comp (f : StructuredArrow S T) (g : f.right ⟶ Y') (g' : Y' ⟶ Y'') : @@ -144,10 +144,10 @@ def mkPostcomp (f : S ⟶ T.obj Y) (g : Y ⟶ Y') : mk f ⟶ mk (f ≫ T.map g) left := 𝟙 _ right := g -lemma mkPostcomp_id (f : S ⟶ T.obj Y) : mkPostcomp f (𝟙 Y) = eqToHom (by aesop_cat) := by aesop_cat +lemma mkPostcomp_id (f : S ⟶ T.obj Y) : mkPostcomp f (𝟙 Y) = eqToHom (by simp) := by simp lemma mkPostcomp_comp (f : S ⟶ T.obj Y) (g : Y ⟶ Y') (g' : Y' ⟶ Y'') : mkPostcomp f (g ≫ g') = mkPostcomp f g ≫ mkPostcomp (f ≫ T.map g) g' ≫ eqToHom (by simp) := by - aesop_cat + simp /-- To construct an isomorphism of structured arrows, we need an isomorphism of the objects underlying the target, @@ -257,7 +257,7 @@ noncomputable def mkIdInitial [T.Full] [T.Faithful] : IsInitial (mk (𝟙 (T.obj desc c := homMk (T.preimage c.pt.hom) uniq c m _ := by apply CommaMorphism.ext - · aesop_cat + · simp · apply T.map_injective simpa only [homMk_right, T.map_preimage, ← w m] using (Category.id_comp _).symm @@ -293,7 +293,7 @@ instance (S : C) (F : B ⥤ C) (G : C ⥤ D) : (post S F G).Faithful where map_injective {_ _} _ _ h := by simpa [ext_iff] using h instance (S : C) (F : B ⥤ C) (G : C ⥤ D) [G.Faithful] : (post S F G).Full where - map_surjective f := ⟨homMk f.right (G.map_injective (by simpa using f.w.symm)), by aesop_cat⟩ + map_surjective f := ⟨homMk f.right (G.map_injective (by simpa using f.w.symm)), by simp⟩ instance (S : C) (F : B ⥤ C) (G : C ⥤ D) [G.Full] : (post S F G).EssSurj where mem_essImage h := ⟨mk (G.preimage h.hom), ⟨isoMk (Iso.refl _) (by simp)⟩⟩ @@ -468,7 +468,7 @@ lemma homMk'_id (f : CostructuredArrow S T) : homMk' f (𝟙 f.left) = eqToHom ( ext simp [eqToHom_left] -lemma homMk'_mk_id (f : S.obj Y ⟶ T) : homMk' (mk f) (𝟙 Y) = eqToHom (by aesop_cat) := +lemma homMk'_mk_id (f : S.obj Y ⟶ T) : homMk' (mk f) (𝟙 Y) = eqToHom (by simp) := homMk'_id _ lemma homMk'_comp (f : CostructuredArrow S T) (g : Y' ⟶ f.left) (g' : Y'' ⟶ Y') : @@ -486,10 +486,10 @@ def mkPrecomp (f : S.obj Y ⟶ T) (g : Y' ⟶ Y) : mk (S.map g ≫ f) ⟶ mk f w left := g right := 𝟙 _ -lemma mkPrecomp_id (f : S.obj Y ⟶ T) : mkPrecomp f (𝟙 Y) = eqToHom (by aesop_cat) := by aesop_cat +lemma mkPrecomp_id (f : S.obj Y ⟶ T) : mkPrecomp f (𝟙 Y) = eqToHom (by simp) := by simp lemma mkPrecomp_comp (f : S.obj Y ⟶ T) (g : Y' ⟶ Y) (g' : Y'' ⟶ Y') : mkPrecomp f (g' ≫ g) = eqToHom (by simp) ≫ mkPrecomp (S.map g ≫ f) g' ≫ mkPrecomp f g := by - aesop_cat + simp /-- To construct an isomorphism of costructured arrows, we need an isomorphism of the objects underlying the source, @@ -634,7 +634,7 @@ instance (F : B ⥤ C) (G : C ⥤ D) (S : C) : (post F G S).Faithful where map_injective {_ _} _ _ h := by simpa [ext_iff] using h instance (F : B ⥤ C) (G : C ⥤ D) (S : C) [G.Faithful] : (post F G S).Full where - map_surjective f := ⟨homMk f.left (G.map_injective (by simpa using f.w)), by aesop_cat⟩ + map_surjective f := ⟨homMk f.left (G.map_injective (by simpa using f.w)), by simp⟩ instance (F : B ⥤ C) (G : C ⥤ D) (S : C) [G.Full] : (post F G S).EssSurj where mem_essImage h := ⟨mk (G.preimage h.hom), ⟨isoMk (Iso.refl _) (by simp)⟩⟩ diff --git a/Mathlib/CategoryTheory/DiscreteCategory.lean b/Mathlib/CategoryTheory/DiscreteCategory.lean index 1d28f0410aa9e..4505a8e1b9584 100644 --- a/Mathlib/CategoryTheory/DiscreteCategory.lean +++ b/Mathlib/CategoryTheory/DiscreteCategory.lean @@ -246,9 +246,9 @@ def equivalence {I : Type u₁} {J : Type u₂} (e : I ≃ J) : Discrete I ≌ D functor := Discrete.functor (Discrete.mk ∘ (e : I → J)) inverse := Discrete.functor (Discrete.mk ∘ (e.symm : J → I)) unitIso := - Discrete.natIso fun i => eqToIso (by aesop_cat) + Discrete.natIso fun i => eqToIso (by simp) counitIso := - Discrete.natIso fun j => eqToIso (by aesop_cat) + Discrete.natIso fun j => eqToIso (by simp) /-- We can convert an equivalence of `discrete` categories to a type-level `Equiv`. -/ @[simps] diff --git a/Mathlib/CategoryTheory/Elements.lean b/Mathlib/CategoryTheory/Elements.lean index 7d85416b30755..01aa8d0ac97b7 100644 --- a/Mathlib/CategoryTheory/Elements.lean +++ b/Mathlib/CategoryTheory/Elements.lean @@ -59,7 +59,7 @@ lemma Functor.Elements.ext {F : C ⥤ Type w} (x y : F.Elements) (h₁ : x.fst = -/ instance categoryOfElements (F : C ⥤ Type w) : Category.{v} F.Elements where Hom p q := { f : p.1 ⟶ q.1 // (F.map f) p.2 = q.2 } - id p := ⟨𝟙 p.1, by aesop_cat⟩ + id p := ⟨𝟙 p.1, by simp⟩ comp {X Y Z} f g := ⟨f.val ≫ g.val, by simp [f.2, g.2]⟩ /-- Natural transformations are mapped to functors between category of elements -/ diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean index 949a01797a9d6..a7842ed6f55b9 100644 --- a/Mathlib/CategoryTheory/Equivalence.lean +++ b/Mathlib/CategoryTheory/Equivalence.lean @@ -292,13 +292,13 @@ def funInvIdAssoc (e : C ≌ D) (F : C ⥤ E) : e.functor ⋙ e.inverse ⋙ F theorem funInvIdAssoc_hom_app (e : C ≌ D) (F : C ⥤ E) (X : C) : (funInvIdAssoc e F).hom.app X = F.map (e.unitInv.app X) := by dsimp [funInvIdAssoc] - aesop_cat + simp @[simp] theorem funInvIdAssoc_inv_app (e : C ≌ D) (F : C ⥤ E) (X : C) : (funInvIdAssoc e F).inv.app X = F.map (e.unit.app X) := by dsimp [funInvIdAssoc] - aesop_cat + simp /-- Composing a functor with both functors of an equivalence yields a naturally isomorphic functor. -/ @@ -309,13 +309,13 @@ def invFunIdAssoc (e : C ≌ D) (F : D ⥤ E) : e.inverse ⋙ e.functor ⋙ F theorem invFunIdAssoc_hom_app (e : C ≌ D) (F : D ⥤ E) (X : D) : (invFunIdAssoc e F).hom.app X = F.map (e.counit.app X) := by dsimp [invFunIdAssoc] - aesop_cat + simp @[simp] theorem invFunIdAssoc_inv_app (e : C ≌ D) (F : D ⥤ E) (X : D) : (invFunIdAssoc e F).inv.app X = F.map (e.counitInv.app X) := by dsimp [invFunIdAssoc] - aesop_cat + simp /-- If `C` is equivalent to `D`, then `C ⥤ E` is equivalent to `D ⥤ E`. -/ @[simps! functor inverse unitIso counitIso] @@ -524,7 +524,7 @@ i.e. faithful, full, and essentially surjective. -/ noncomputable def inv (F : C ⥤ D) [F.IsEquivalence] : D ⥤ C where obj X := F.objPreimage X map {X Y} f := F.preimage ((F.objObjPreimageIso X).hom ≫ f ≫ (F.objObjPreimageIso Y).inv) - map_id X := by apply F.map_injective; aesop_cat + map_id X := by apply F.map_injective; simp map_comp {X Y Z} f g := by apply F.map_injective; simp /-- Interpret a functor that is an equivalence as an equivalence. diff --git a/Mathlib/CategoryTheory/Functor/Const.lean b/Mathlib/CategoryTheory/Functor/Const.lean index ba2d6a7b37e98..63504ca0530c7 100644 --- a/Mathlib/CategoryTheory/Functor/Const.lean +++ b/Mathlib/CategoryTheory/Functor/Const.lean @@ -96,7 +96,7 @@ instance [Nonempty J] : Faithful (const J : C ⥤ J ⥤ C) where def compConstIso (F : C ⥤ D) : F ⋙ Functor.const J ≅ Functor.const J ⋙ (whiskeringRight J C D).obj F := NatIso.ofComponents - (fun X => NatIso.ofComponents (fun _ => Iso.refl _) (by aesop_cat)) + (fun X => NatIso.ofComponents (fun _ => Iso.refl _) (by simp)) (by aesop_cat) /-- The canonical isomorphism diff --git a/Mathlib/CategoryTheory/Functor/Derived/RightDerived.lean b/Mathlib/CategoryTheory/Functor/Derived/RightDerived.lean index 8acc96f8e2d2d..c815ff3015b16 100644 --- a/Mathlib/CategoryTheory/Functor/Derived/RightDerived.lean +++ b/Mathlib/CategoryTheory/Functor/Derived/RightDerived.lean @@ -117,7 +117,7 @@ lemma rightDerivedNatTrans_app (τ : F ⟶ F') (X : C) : @[simp] lemma rightDerivedNatTrans_id : rightDerivedNatTrans RF RF α α W (𝟙 F) = 𝟙 RF := - rightDerived_ext RF α W _ _ _ (by aesop_cat) + rightDerived_ext RF α W _ _ _ (by simp) variable [RF'.IsRightDerivedFunctor α' W] @@ -125,7 +125,7 @@ variable [RF'.IsRightDerivedFunctor α' W] lemma rightDerivedNatTrans_comp (τ : F ⟶ F') (τ' : F' ⟶ F'') : rightDerivedNatTrans RF RF' α α' W τ ≫ rightDerivedNatTrans RF' RF'' α' α'' W τ' = rightDerivedNatTrans RF RF'' α α'' W (τ ≫ τ') := - rightDerived_ext RF α W _ _ _ (by aesop_cat) + rightDerived_ext RF α W _ _ _ (by simp) /-- The natural isomorphism `RF ≅ RF'` on right derived functors that is induced by a natural isomorphism `F ≅ F'`. -/ diff --git a/Mathlib/CategoryTheory/Functor/FunctorHom.lean b/Mathlib/CategoryTheory/Functor/FunctorHom.lean index 3f4312c687eaa..be91d1395df1e 100644 --- a/Mathlib/CategoryTheory/Functor/FunctorHom.lean +++ b/Mathlib/CategoryTheory/Functor/FunctorHom.lean @@ -131,7 +131,7 @@ def functorHomEquiv (A : C ⥤ Type max u v v') : (A ⟶ F.functorHom G) ≃ Hom ext X a Y f exact (HomObj.congr_app (congr_fun (φ.naturality f) a) Y (𝟙 _)).trans (congr_arg ((φ.app X a).app Y) (by simp)) - right_inv x := by aesop + right_inv x := by simp variable {F G} in /-- Morphisms `(𝟙_ (C ⥤ Type max v' v u) ⟶ F.functorHom G)` are in bijection with diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean index 4ea68229bb26b..b969b1f888ff4 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean @@ -102,7 +102,7 @@ noncomputable def homEquivOfIsRightKanExtension (G : D ⥤ H) : (G ⟶ F') ≃ (L ⋙ G ⟶ F) where toFun β := whiskerLeft _ β ≫ α invFun β := liftOfIsRightKanExtension _ α _ β - left_inv β := Functor.hom_ext_of_isRightKanExtension _ α _ _ (by aesop_cat) + left_inv β := Functor.hom_ext_of_isRightKanExtension _ α _ _ (by simp) right_inv := by aesop_cat lemma isRightKanExtension_of_iso {F' F'' : D ⥤ H} (e : F' ≅ F'') {L : C ⥤ D} {F : C ⥤ H} @@ -193,7 +193,7 @@ noncomputable def homEquivOfIsLeftKanExtension (G : D ⥤ H) : (F' ⟶ G) ≃ (F ⟶ L ⋙ G) where toFun β := α ≫ whiskerLeft _ β invFun β := descOfIsLeftKanExtension _ α _ β - left_inv β := Functor.hom_ext_of_isLeftKanExtension _ α _ _ (by aesop_cat) + left_inv β := Functor.hom_ext_of_isLeftKanExtension _ α _ _ (by simp) right_inv := by aesop_cat lemma isLeftKanExtension_of_iso {F' : D ⥤ H} {F'' : D ⥤ H} (e : F' ≅ F'') diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean index a1a044d1a3846..12c6638bcdfe2 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean @@ -140,8 +140,8 @@ def isPointwiseLeftKanExtensionEquivOfIso (e : E ≅ E') : E.IsPointwiseLeftKanExtension ≃ E'.IsPointwiseLeftKanExtension where toFun h := fun Y => (isPointwiseLeftKanExtensionAtEquivOfIso e Y) (h Y) invFun h := fun Y => (isPointwiseLeftKanExtensionAtEquivOfIso e Y).symm (h Y) - left_inv h := by aesop - right_inv h := by aesop + left_inv h := by simp + right_inv h := by simp variable (h : E.IsPointwiseLeftKanExtension) include h @@ -277,8 +277,8 @@ def isPointwiseRightKanExtensionEquivOfIso (e : E ≅ E') : E.IsPointwiseRightKanExtension ≃ E'.IsPointwiseRightKanExtension where toFun h := fun Y => (isPointwiseRightKanExtensionAtEquivOfIso e Y) (h Y) invFun h := fun Y => (isPointwiseRightKanExtensionAtEquivOfIso e Y).symm (h Y) - left_inv h := by aesop - right_inv h := by aesop + left_inv h := by simp + right_inv h := by simp variable (h : E.IsPointwiseRightKanExtension) include h diff --git a/Mathlib/CategoryTheory/Galois/Decomposition.lean b/Mathlib/CategoryTheory/Galois/Decomposition.lean index 0266fabf8bc16..37f5ba335107f 100644 --- a/Mathlib/CategoryTheory/Galois/Decomposition.lean +++ b/Mathlib/CategoryTheory/Galois/Decomposition.lean @@ -66,7 +66,7 @@ private lemma has_decomp_connected_components_aux_initial (X : C) (h : IsInitial ∃ (ι : Type) (f : ι → C) (g : (i : ι) → (f i) ⟶ X) (_ : IsColimit (Cofan.mk X g)), (∀ i, IsConnected (f i)) ∧ Finite ι := by refine ⟨Empty, fun _ ↦ X, fun _ ↦ 𝟙 X, ?_⟩ - use mkCofanColimit _ (fun s ↦ IsInitial.to h s.pt) (fun s ↦ by aesop) + use mkCofanColimit _ (fun s ↦ IsInitial.to h s.pt) (fun s ↦ by simp) (fun s m _ ↦ IsInitial.hom_ext h m _) exact ⟨by simp only [IsEmpty.forall_iff], inferInstance⟩ diff --git a/Mathlib/CategoryTheory/GlueData.lean b/Mathlib/CategoryTheory/GlueData.lean index 2c0b603d6f43f..fcea852442c19 100644 --- a/Mathlib/CategoryTheory/GlueData.lean +++ b/Mathlib/CategoryTheory/GlueData.lean @@ -443,9 +443,9 @@ def GlueData'.t'' (D : GlueData' C) (i j k : D.J) : else haveI := Ne.symm hij pullback.map _ _ _ _ (eqToHom (by aesop)) (eqToHom (by rw [dif_neg hik])) - (eqToHom (by aesop)) (by delta f'; aesop) (by delta f'; aesop) ≫ + (eqToHom (by simp)) (by delta f'; aesop) (by delta f'; aesop) ≫ D.t' i j k hij hik hjk ≫ - pullback.map _ _ _ _ (eqToHom (by aesop)) (eqToHom (by aesop)) (eqToHom (by aesop)) + pullback.map _ _ _ _ (eqToHom (by aesop)) (eqToHom (by aesop)) (eqToHom (by simp)) (by delta f'; aesop) (by delta f'; aesop) open scoped Classical in diff --git a/Mathlib/CategoryTheory/GradedObject/Single.lean b/Mathlib/CategoryTheory/GradedObject/Single.lean index 180a200b1ed92..e5add88c366ef 100644 --- a/Mathlib/CategoryTheory/GradedObject/Single.lean +++ b/Mathlib/CategoryTheory/GradedObject/Single.lean @@ -77,7 +77,7 @@ variable (C) in evaluation functor `eval j` identifies to the identity functor. -/ @[simps!] noncomputable def singleCompEval (j : J) : single j ⋙ eval j ≅ 𝟭 C := - NatIso.ofComponents (singleObjApplyIso j) (by aesop_cat) + NatIso.ofComponents (singleObjApplyIso j) (by simp) end GradedObject diff --git a/Mathlib/CategoryTheory/Groupoid.lean b/Mathlib/CategoryTheory/Groupoid.lean index a880d7c21b5a3..5b43b7b7a2a4b 100644 --- a/Mathlib/CategoryTheory/Groupoid.lean +++ b/Mathlib/CategoryTheory/Groupoid.lean @@ -95,7 +95,7 @@ variable (X Y) /-- In a groupoid, isomorphisms are equivalent to morphisms. -/ def Groupoid.isoEquivHom : (X ≅ Y) ≃ (X ⟶ Y) where toFun := Iso.hom - invFun f := ⟨f, Groupoid.inv f, (by aesop_cat), (by aesop_cat)⟩ + invFun f := ⟨f, Groupoid.inv f, (by simp), (by simp)⟩ left_inv _ := Iso.ext rfl right_inv _ := rfl diff --git a/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean b/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean index e227d08f5cbc7..8c6c94626e7a8 100644 --- a/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean +++ b/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean @@ -103,7 +103,7 @@ def functorExtension₁CompWhiskeringLeftToKaroubiIso : (fun X => { hom := { f := (F.obj X).p } inv := { f := (F.obj X).p } }) - (fun {X Y} f => by aesop_cat)) + (fun {X Y} f => by simp)) (by aesop_cat) /-- The counit isomorphism of the equivalence `(C ⥤ Karoubi D) ≌ (Karoubi C ⥤ Karoubi D)`. -/ @@ -180,7 +180,7 @@ def functorExtension₂CompWhiskeringLeftToKaroubiIso : (fun X => { hom := { f := 𝟙 _ } inv := { f := 𝟙 _ } }) - (by aesop_cat)) + (by simp)) (by aesop_cat) section IsIdempotentComplete diff --git a/Mathlib/CategoryTheory/Idempotents/HomologicalComplex.lean b/Mathlib/CategoryTheory/Idempotents/HomologicalComplex.lean index 52ef588b3f18e..3efddee1016cd 100644 --- a/Mathlib/CategoryTheory/Idempotents/HomologicalComplex.lean +++ b/Mathlib/CategoryTheory/Idempotents/HomologicalComplex.lean @@ -120,7 +120,7 @@ def inverse : HomologicalComplex (Karoubi C) c ⥤ Karoubi (HomologicalComplex C `Karoubi (HomologicalComplex C c) ≌ HomologicalComplex (Karoubi C) c`. -/ @[simps!] def counitIso : inverse ⋙ functor ≅ 𝟭 (HomologicalComplex (Karoubi C) c) := - eqToIso (Functor.ext (fun P => HomologicalComplex.ext (by aesop_cat) (by aesop_cat)) + eqToIso (Functor.ext (fun P => HomologicalComplex.ext (by aesop_cat) (by simp)) (by aesop_cat)) /-- The unit isomorphism of the equivalence diff --git a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean index d4b4fc8a4f69e..aa8ebd6d96a7d 100644 --- a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean @@ -273,11 +273,11 @@ theorem decompId_p_toKaroubi (X : C) : decompId_p ((toKaroubi C).obj X) = 𝟙 _ theorem decompId_i_naturality {P Q : Karoubi C} (f : P ⟶ Q) : f ≫ decompId_i Q = decompId_i P ≫ (by exact Hom.mk f.f (by simp)) := by - aesop_cat + simp theorem decompId_p_naturality {P Q : Karoubi C} (f : P ⟶ Q) : decompId_p P ≫ f = (by exact Hom.mk f.f (by simp)) ≫ decompId_p Q := by - aesop_cat + simp @[simp] theorem zsmul_hom [Preadditive C] {P Q : Karoubi C} (n : ℤ) (f : P ⟶ Q) : (n • f).f = n • f.f := diff --git a/Mathlib/CategoryTheory/Idempotents/KaroubiKaroubi.lean b/Mathlib/CategoryTheory/Idempotents/KaroubiKaroubi.lean index 2fc3351f029a2..4195b5218d23f 100644 --- a/Mathlib/CategoryTheory/Idempotents/KaroubiKaroubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/KaroubiKaroubi.lean @@ -46,7 +46,7 @@ instance [Preadditive C] : Functor.Additive (inverse C) where /-- The unit isomorphism of the equivalence -/ @[simps!] def unitIso : 𝟭 (Karoubi C) ≅ toKaroubi (Karoubi C) ⋙ inverse C := - eqToIso (Functor.ext (by aesop_cat) (by aesop_cat)) + eqToIso (Functor.ext (by aesop_cat) (by simp)) attribute [local simp] p_comm_f in /-- The counit isomorphism of the equivalence -/ diff --git a/Mathlib/CategoryTheory/IsConnected.lean b/Mathlib/CategoryTheory/IsConnected.lean index 73adc542b3de7..9f27285ea78dc 100644 --- a/Mathlib/CategoryTheory/IsConnected.lean +++ b/Mathlib/CategoryTheory/IsConnected.lean @@ -102,7 +102,7 @@ def isoConstant [IsPreconnected J] {α : Type u₂} (F : J ⥤ Discrete α) (j : F ≅ (Functor.const J).obj (F.obj j) := (IsPreconnected.IsoConstantAux.factorThroughDiscrete F).symm ≪≫ isoWhiskerRight (IsPreconnected.iso_constant _ j).some _ - ≪≫ NatIso.ofComponents (fun _ => eqToIso Function.apply_invFun_apply) (by aesop_cat) + ≪≫ NatIso.ofComponents (fun _ => eqToIso Function.apply_invFun_apply) (by simp) /-- If `J` is connected, any functor to a discrete category is constant on objects. The converse is given in `IsConnected.of_any_functor_const_on_obj`. diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index 6fee46c53e64c..68bb5355b5bfe 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -158,7 +158,7 @@ instance inhabitedCocone (F : Discrete PUnit ⥤ C) : Inhabited (Cocone F) := intro X Y f match X, Y, f with | .mk A, .mk B, .up g => - aesop_cat + simp } }⟩ @@ -408,7 +408,7 @@ def functoriality : Cone F ⥤ Cone (F ⋙ G) where { pt := G.obj A.pt π := { app := fun j => G.map (A.π.app j) - naturality := by intros; erw [← G.map_comp]; aesop_cat } } + naturality := by intros; erw [← G.map_comp]; simp } } map f := { hom := G.map f.hom w := fun j => by simp [-ConeMorphism.w, ← f.w j] } @@ -606,7 +606,7 @@ def functoriality : Cocone F ⥤ Cocone (F ⋙ G) where { pt := G.obj A.pt ι := { app := fun j => G.map (A.ι.app j) - naturality := by intros; erw [← G.map_comp]; aesop_cat } } + naturality := by intros; erw [← G.map_comp]; simp } } map f := { hom := G.map f.hom w := by intros; rw [← Functor.map_comp, CoconeMorphism.w] } diff --git a/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean b/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean index 5712931af86a2..12841a497aa58 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean @@ -68,7 +68,7 @@ def liftToFinsetColimitCocone [HasColimitsOfShape (Finset (Discrete α)) C] convert h j using 1 · simp [← colimit.w (liftToFinsetObj F) ⟨⟨Finset.singleton_subset_iff.2 hj⟩⟩] rfl - · aesop_cat } + · simp } variable (C) (α) @@ -180,7 +180,7 @@ def liftToFinsetLimitCone [HasLimitsOfShape (Finset (Discrete α))ᵒᵖ C] convert h j using 1 · simp [← limit.w (liftToFinsetObj F) ⟨⟨⟨Finset.singleton_subset_iff.2 hj⟩⟩⟩] rfl - · aesop_cat } + · simp } variable (C) (α) diff --git a/Mathlib/CategoryTheory/Limits/Constructions/ZeroObjects.lean b/Mathlib/CategoryTheory/Limits/Constructions/ZeroObjects.lean index d978a5c54dda2..e9e831e516ae5 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/ZeroObjects.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/ZeroObjects.lean @@ -32,7 +32,7 @@ def binaryFanZeroLeft (X : C) : BinaryFan (0 : C) X := /-- The limit cone for the product with a zero object is limiting. -/ def binaryFanZeroLeftIsLimit (X : C) : IsLimit (binaryFanZeroLeft X) := - BinaryFan.isLimitMk (fun s => BinaryFan.snd s) (by aesop_cat) (by aesop_cat) + BinaryFan.isLimitMk (fun s => BinaryFan.snd s) (by aesop_cat) (by simp) (fun s m _ h₂ => by simpa using h₂) instance hasBinaryProduct_zero_left (X : C) : HasBinaryProduct (0 : C) X := @@ -57,7 +57,7 @@ def binaryFanZeroRight (X : C) : BinaryFan X (0 : C) := /-- The limit cone for the product with a zero object is limiting. -/ def binaryFanZeroRightIsLimit (X : C) : IsLimit (binaryFanZeroRight X) := - BinaryFan.isLimitMk (fun s => BinaryFan.fst s) (by aesop_cat) (by aesop_cat) + BinaryFan.isLimitMk (fun s => BinaryFan.fst s) (by simp) (by aesop_cat) (fun s m h₁ _ => by simpa using h₁) instance hasBinaryProduct_zero_right (X : C) : HasBinaryProduct X (0 : C) := @@ -82,7 +82,7 @@ def binaryCofanZeroLeft (X : C) : BinaryCofan (0 : C) X := /-- The colimit cocone for the coproduct with a zero object is colimiting. -/ def binaryCofanZeroLeftIsColimit (X : C) : IsColimit (binaryCofanZeroLeft X) := - BinaryCofan.isColimitMk (fun s => BinaryCofan.inr s) (by aesop_cat) (by aesop_cat) + BinaryCofan.isColimitMk (fun s => BinaryCofan.inr s) (by aesop_cat) (by simp) (fun s m _ h₂ => by simpa using h₂) instance hasBinaryCoproduct_zero_left (X : C) : HasBinaryCoproduct (0 : C) X := @@ -107,7 +107,7 @@ def binaryCofanZeroRight (X : C) : BinaryCofan X (0 : C) := /-- The colimit cocone for the coproduct with a zero object is colimiting. -/ def binaryCofanZeroRightIsColimit (X : C) : IsColimit (binaryCofanZeroRight X) := - BinaryCofan.isColimitMk (fun s => BinaryCofan.inl s) (by aesop_cat) (by aesop_cat) + BinaryCofan.isColimitMk (fun s => BinaryCofan.inl s) (by simp) (by aesop_cat) (fun s m h₁ _ => by simpa using h₁) instance hasBinaryCoproduct_zero_right (X : C) : HasBinaryCoproduct X (0 : C) := diff --git a/Mathlib/CategoryTheory/Limits/HasLimits.lean b/Mathlib/CategoryTheory/Limits/HasLimits.lean index 27bcb80063f88..c3297674c5db3 100644 --- a/Mathlib/CategoryTheory/Limits/HasLimits.lean +++ b/Mathlib/CategoryTheory/Limits/HasLimits.lean @@ -225,13 +225,13 @@ def limit.isoLimitCone {F : J ⥤ C} [HasLimit F] (t : LimitCone F) : limit F theorem limit.isoLimitCone_hom_π {F : J ⥤ C} [HasLimit F] (t : LimitCone F) (j : J) : (limit.isoLimitCone t).hom ≫ t.cone.π.app j = limit.π F j := by dsimp [limit.isoLimitCone, IsLimit.conePointUniqueUpToIso] - aesop_cat + simp @[reassoc (attr := simp)] theorem limit.isoLimitCone_inv_π {F : J ⥤ C} [HasLimit F] (t : LimitCone F) (j : J) : (limit.isoLimitCone t).inv ≫ limit.π F j = t.cone.π.app j := by dsimp [limit.isoLimitCone, IsLimit.conePointUniqueUpToIso] - aesop_cat + simp @[ext] theorem limit.hom_ext {F : J ⥤ C} [HasLimit F] {X : C} {f f' : X ⟶ limit F} @@ -750,13 +750,13 @@ def colimit.isoColimitCocone {F : J ⥤ C} [HasColimit F] (t : ColimitCocone F) theorem colimit.isoColimitCocone_ι_hom {F : J ⥤ C} [HasColimit F] (t : ColimitCocone F) (j : J) : colimit.ι F j ≫ (colimit.isoColimitCocone t).hom = t.cocone.ι.app j := by dsimp [colimit.isoColimitCocone, IsColimit.coconePointUniqueUpToIso] - aesop_cat + simp @[reassoc (attr := simp)] theorem colimit.isoColimitCocone_ι_inv {F : J ⥤ C} [HasColimit F] (t : ColimitCocone F) (j : J) : t.cocone.ι.app j ≫ (colimit.isoColimitCocone t).inv = colimit.ι F j := by dsimp [colimit.isoColimitCocone, IsColimit.coconePointUniqueUpToIso] - aesop_cat + simp @[ext] theorem colimit.hom_ext {F : J ⥤ C} [HasColimit F] {X : C} {f f' : colimit F ⟶ X} diff --git a/Mathlib/CategoryTheory/Limits/IsLimit.lean b/Mathlib/CategoryTheory/Limits/IsLimit.lean index d0aebd2e0120a..6c1343fba3041 100644 --- a/Mathlib/CategoryTheory/Limits/IsLimit.lean +++ b/Mathlib/CategoryTheory/Limits/IsLimit.lean @@ -263,8 +263,8 @@ def conePointsIsoOfNatIso {F G : J ⥤ C} {s : Cone F} {t : Cone G} (P : IsLimit (w : F ≅ G) : s.pt ≅ t.pt where hom := Q.map s w.hom inv := P.map t w.inv - hom_inv_id := P.hom_ext (by aesop_cat) - inv_hom_id := Q.hom_ext (by aesop_cat) + hom_inv_id := P.hom_ext (by simp) + inv_hom_id := Q.hom_ext (by simp) @[reassoc] theorem conePointsIsoOfNatIso_hom_comp {F G : J ⥤ C} {s : Cone F} {t : Cone G} (P : IsLimit s) @@ -723,8 +723,8 @@ def coconePointsIsoOfNatIso {F G : J ⥤ C} {s : Cocone F} {t : Cocone G} (P : I (Q : IsColimit t) (w : F ≅ G) : s.pt ≅ t.pt where hom := P.map t w.hom inv := Q.map s w.inv - hom_inv_id := P.hom_ext (by aesop_cat) - inv_hom_id := Q.hom_ext (by aesop_cat) + hom_inv_id := P.hom_ext (by simp) + inv_hom_id := Q.hom_ext (by simp) @[reassoc] theorem comp_coconePointsIsoOfNatIso_hom {F G : J ⥤ C} {s : Cocone F} {t : Cocone G} diff --git a/Mathlib/CategoryTheory/Limits/MonoCoprod.lean b/Mathlib/CategoryTheory/Limits/MonoCoprod.lean index 9922662f9f26a..868732f7d96d1 100644 --- a/Mathlib/CategoryTheory/Limits/MonoCoprod.lean +++ b/Mathlib/CategoryTheory/Limits/MonoCoprod.lean @@ -60,7 +60,7 @@ theorem binaryCofan_inr {A B : C} [MonoCoprod C] (c : BinaryCofan A B) (hc : IsC Mono c.inr := by haveI hc' : IsColimit (BinaryCofan.mk c.inr c.inl) := BinaryCofan.IsColimit.mk _ (fun f₁ f₂ => hc.desc (BinaryCofan.mk f₂ f₁)) - (by aesop_cat) (by aesop_cat) + (by simp) (by simp) (fun f₁ f₂ m h₁ h₂ => BinaryCofan.IsColimit.hom_ext hc (by aesop_cat) (by aesop_cat)) exact binaryCofan_inl _ hc' diff --git a/Mathlib/CategoryTheory/Limits/Opposites.lean b/Mathlib/CategoryTheory/Limits/Opposites.lean index 6c32f2d694aba..76a2af3431f75 100644 --- a/Mathlib/CategoryTheory/Limits/Opposites.lean +++ b/Mathlib/CategoryTheory/Limits/Opposites.lean @@ -888,7 +888,7 @@ theorem unop_fst {X Y Z : Cᵒᵖ} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocon c.unop.fst = c.inl.unop := by simp theorem unop_snd {X Y Z : Cᵒᵖ} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : - c.unop.snd = c.inr.unop := by aesop_cat + c.unop.snd = c.inr.unop := by simp -- Porting note: it was originally @[simps (config := lemmasOnly)] /-- The obvious map `PushoutCocone f.op g.op → PullbackCone f g` -/ @@ -898,10 +898,10 @@ def op {X Y Z : C} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : Pullbac (Cone.whisker walkingSpanOpEquiv.inverse (Cocone.op c)) theorem op_fst {X Y Z : C} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : - c.op.fst = c.inl.op := by aesop_cat + c.op.fst = c.inl.op := by simp theorem op_snd {X Y Z : C} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : - c.op.snd = c.inr.op := by aesop_cat + c.op.snd = c.inr.op := by simp end PushoutCocone @@ -917,10 +917,10 @@ def unop {X Y Z : Cᵒᵖ} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : (Cone.whisker walkingSpanOpEquiv.functor c)) theorem unop_inl {X Y Z : Cᵒᵖ} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : - c.unop.inl = c.fst.unop := by aesop_cat + c.unop.inl = c.fst.unop := by simp theorem unop_inr {X Y Z : Cᵒᵖ} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : - c.unop.inr = c.snd.unop := by aesop_cat + c.unop.inr = c.snd.unop := by simp /-- The obvious map `PullbackCone f g → PushoutCocone f.op g.op` -/ @[simps!] @@ -929,10 +929,10 @@ def op {X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : PushoutC (Cocone.whisker walkingCospanOpEquiv.inverse (Cone.op c)) theorem op_inl {X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : - c.op.inl = c.fst.op := by aesop_cat + c.op.inl = c.fst.op := by simp theorem op_inr {X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : - c.op.inr = c.snd.op := by aesop_cat + c.op.inr = c.snd.op := by simp /-- If `c` is a pullback cone, then `c.op.unop` is isomorphic to `c`. -/ def opUnop {X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) : c.op.unop ≅ c := diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean index 550f8c8bea18d..95581e7fd2454 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean @@ -237,7 +237,7 @@ theorem biproductComparison'_comp_biproductComparison : @[simps] def splitEpiBiproductComparison : SplitEpi (biproductComparison F f) where section_ := biproductComparison' F f - id := by aesop + id := by simp instance : IsSplitEpi (biproductComparison F f) := IsSplitEpi.mk' (splitEpiBiproductComparison F f) @@ -246,7 +246,7 @@ instance : IsSplitEpi (biproductComparison F f) := @[simps] def splitMonoBiproductComparison' : SplitMono (biproductComparison' F f) where retraction := biproductComparison F f - id := by aesop + id := by simp instance : IsSplitMono (biproductComparison' F f) := IsSplitMono.mk' (splitMonoBiproductComparison' F f) @@ -321,7 +321,7 @@ theorem biprodComparison'_comp_biprodComparison : @[simps] def splitEpiBiprodComparison : SplitEpi (biprodComparison F X Y) where section_ := biprodComparison' F X Y - id := by aesop + id := by simp instance : IsSplitEpi (biprodComparison F X Y) := IsSplitEpi.mk' (splitEpiBiprodComparison F X Y) @@ -330,7 +330,7 @@ instance : IsSplitEpi (biprodComparison F X Y) := @[simps] def splitMonoBiprodComparison' : SplitMono (biprodComparison' F X Y) where retraction := biprodComparison F X Y - id := by aesop + id := by simp instance : IsSplitMono (biprodComparison' F X Y) := IsSplitMono.mk' (splitMonoBiprodComparison' F X Y) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index cf91656226688..28ea0694e2027 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -280,7 +280,7 @@ def whiskerToCone {f : J → C} (c : Bicone f) (g : K ≃ J) : (c.whisker g).toCone ≅ (Cones.postcompose (Discrete.functorComp f g).inv).obj (c.toCone.whisker (Discrete.functor (Discrete.mk ∘ g))) := - Cones.ext (Iso.refl _) (by aesop_cat) + Cones.ext (Iso.refl _) (by simp) /-- Taking the cocone of a whiskered bicone results in a cone isomorphic to one gained by whiskering the cocone and precomposing with a suitable isomorphism. -/ @@ -288,7 +288,7 @@ def whiskerToCocone {f : J → C} (c : Bicone f) (g : K ≃ J) : (c.whisker g).toCocone ≅ (Cocones.precompose (Discrete.functorComp f g).hom).obj (c.toCocone.whisker (Discrete.functor (Discrete.mk ∘ g))) := - Cocones.ext (Iso.refl _) (by aesop_cat) + Cocones.ext (Iso.refl _) (by simp) /-- Whiskering a bicone with an equivalence between types preserves being a bilimit bicone. -/ noncomputable def whiskerIsBilimitIff {f : J → C} (c : Bicone f) (g : K ≃ J) : @@ -625,7 +625,7 @@ instance biproduct.map_epi {f g : J → C} [HasBiproduct f] [HasBiproduct g] (p ι_colimMap_assoc, Discrete.functor_obj_eq_as, Discrete.natTrans_app, colimit.ι_desc_assoc, Cofan.mk_pt, Cofan.mk_ι_app, ι_π, ι_π_assoc] split - all_goals aesop + all_goals simp_all rw [this] infer_instance @@ -1757,7 +1757,7 @@ instance biprod.map_epi {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [Epi f] instance prod.map_epi {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [Epi f] [Epi g] [HasBinaryBiproduct W X] [HasBinaryBiproduct Y Z] : Epi (prod.map f g) := by rw [show prod.map f g = (biprod.isoProd _ _).inv ≫ biprod.map f g ≫ - (biprod.isoProd _ _).hom by aesop] + (biprod.isoProd _ _).hom by simp] infer_instance instance biprod.map_mono {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [Mono f] @@ -1769,7 +1769,7 @@ instance biprod.map_mono {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [Mono f] instance coprod.map_mono {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [Mono f] [Mono g] [HasBinaryBiproduct W X] [HasBinaryBiproduct Y Z] : Mono (coprod.map f g) := by rw [show coprod.map f g = (biprod.isoCoprod _ _).inv ≫ biprod.map f g ≫ - (biprod.isoCoprod _ _).hom by aesop] + (biprod.isoCoprod _ _).hom by simp] infer_instance section BiprodKernel diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean b/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean index 36a0a6e065433..fdf4259edc924 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean @@ -533,7 +533,7 @@ theorem Cofork.IsColimit.homIso_natural {X Y : C} {f g : X ⟶ Y} {t : Cofork f def Cone.ofFork {F : WalkingParallelPair ⥤ C} (t : Fork (F.map left) (F.map right)) : Cone F where pt := t.pt π := - { app := fun X => t.π.app X ≫ eqToHom (by aesop) + { app := fun X => t.π.app X ≫ eqToHom (by simp) naturality := by rintro _ _ (_|_|_) <;> {dsimp; simp [t.condition]}} /-- This is a helper construction that can be useful when verifying that a category has all @@ -548,23 +548,23 @@ def Cocone.ofCofork {F : WalkingParallelPair ⥤ C} (t : Cofork (F.map left) (F. Cocone F where pt := t.pt ι := - { app := fun X => eqToHom (by aesop) ≫ t.ι.app X + { app := fun X => eqToHom (by simp) ≫ t.ι.app X naturality := by rintro _ _ (_|_|_) <;> {dsimp; simp [t.condition]}} @[simp] theorem Cone.ofFork_π {F : WalkingParallelPair ⥤ C} (t : Fork (F.map left) (F.map right)) (j) : - (Cone.ofFork t).π.app j = t.π.app j ≫ eqToHom (by aesop) := rfl + (Cone.ofFork t).π.app j = t.π.app j ≫ eqToHom (by simp) := rfl @[simp] theorem Cocone.ofCofork_ι {F : WalkingParallelPair ⥤ C} (t : Cofork (F.map left) (F.map right)) - (j) : (Cocone.ofCofork t).ι.app j = eqToHom (by aesop) ≫ t.ι.app j := rfl + (j) : (Cocone.ofCofork t).ι.app j = eqToHom (by simp) ≫ t.ι.app j := rfl /-- Given `F : WalkingParallelPair ⥤ C`, which is really the same as `parallelPair (F.map left) (F.map right)` and a cone on `F`, we get a fork on `F.map left` and `F.map right`. -/ def Fork.ofCone {F : WalkingParallelPair ⥤ C} (t : Cone F) : Fork (F.map left) (F.map right) where pt := t.pt - π := { app := fun X => t.π.app X ≫ eqToHom (by aesop) + π := { app := fun X => t.π.app X ≫ eqToHom (by simp) naturality := by rintro _ _ (_|_|_) <;> {dsimp; simp}} /-- Given `F : WalkingParallelPair ⥤ C`, which is really the same as @@ -573,16 +573,16 @@ def Fork.ofCone {F : WalkingParallelPair ⥤ C} (t : Cone F) : Fork (F.map left) def Cofork.ofCocone {F : WalkingParallelPair ⥤ C} (t : Cocone F) : Cofork (F.map left) (F.map right) where pt := t.pt - ι := { app := fun X => eqToHom (by aesop) ≫ t.ι.app X + ι := { app := fun X => eqToHom (by simp) ≫ t.ι.app X naturality := by rintro _ _ (_|_|_) <;> {dsimp; simp}} @[simp] theorem Fork.ofCone_π {F : WalkingParallelPair ⥤ C} (t : Cone F) (j) : - (Fork.ofCone t).π.app j = t.π.app j ≫ eqToHom (by aesop) := rfl + (Fork.ofCone t).π.app j = t.π.app j ≫ eqToHom (by simp) := rfl @[simp] theorem Cofork.ofCocone_ι {F : WalkingParallelPair ⥤ C} (t : Cocone F) (j) : - (Cofork.ofCocone t).ι.app j = eqToHom (by aesop) ≫ t.ι.app j := rfl + (Cofork.ofCocone t).ι.app j = eqToHom (by simp) ≫ t.ι.app j := rfl @[simp] theorem Fork.ι_postcompose {f' g' : X ⟶ Y} {α : parallelPair f g ⟶ parallelPair f' g'} @@ -710,7 +710,7 @@ theorem equalizer.condition : equalizer.ι f g ≫ f = equalizer.ι f g ≫ g := /-- The equalizer built from `equalizer.ι f g` is limiting. -/ noncomputable def equalizerIsEqualizer : IsLimit (Fork.ofι (equalizer.ι f g) (equalizer.condition f g)) := - IsLimit.ofIsoLimit (limit.isLimit _) (Fork.ext (Iso.refl _) (by aesop)) + IsLimit.ofIsoLimit (limit.isLimit _) (Fork.ext (Iso.refl _) (by simp)) variable {f g} @@ -859,7 +859,7 @@ theorem coequalizer.condition : f ≫ coequalizer.π f g = g ≫ coequalizer.π /-- The cofork built from `coequalizer.π f g` is colimiting. -/ noncomputable def coequalizerIsCoequalizer : IsColimit (Cofork.ofπ (coequalizer.π f g) (coequalizer.condition f g)) := - IsColimit.ofIsoColimit (colimit.isColimit _) (Cofork.ext (Iso.refl _) (by aesop)) + IsColimit.ofIsoColimit (colimit.isColimit _) (Cofork.ext (Iso.refl _) (by simp)) variable {f g} diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean index 1214bfe40be59..f11792f5eaace 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean @@ -613,7 +613,7 @@ attribute [inherit_doc ImageMap] ImageMap.map ImageMap.map_ι attribute [-simp, nolint simpNF] ImageMap.mk.injEq instance inhabitedImageMap {f : Arrow C} [HasImage f.hom] : Inhabited (ImageMap (𝟙 f)) := - ⟨⟨𝟙 _, by aesop⟩⟩ + ⟨⟨𝟙 _, by simp⟩⟩ attribute [reassoc (attr := simp)] ImageMap.map_ι diff --git a/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean b/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean index de6ee324b9d3d..a33d65333e024 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean @@ -57,11 +57,11 @@ abbrev IsInitial (X : C) := /-- An object `Y` is terminal iff for every `X` there is a unique morphism `X ⟶ Y`. -/ def isTerminalEquivUnique (F : Discrete.{0} PEmpty.{1} ⥤ C) (Y : C) : - IsLimit (⟨Y, by aesop_cat, by aesop_cat⟩ : Cone F) ≃ ∀ X : C, Unique (X ⟶ Y) where + IsLimit (⟨Y, by aesop_cat, by simp⟩ : Cone F) ≃ ∀ X : C, Unique (X ⟶ Y) where toFun t X := - { default := t.lift ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ + { default := t.lift ⟨X, ⟨by aesop_cat, by simp⟩⟩ uniq := fun f => - t.uniq ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ f (by aesop_cat) } + t.uniq ⟨X, ⟨by aesop_cat, by simp⟩⟩ f (by simp) } invFun u := { lift := fun s => (u s.pt).default uniq := fun s _ _ => (u s.pt).2 _ } @@ -103,10 +103,10 @@ def IsTerminal.equivOfIso {X Y : C} (e : X ≅ Y) : /-- An object `X` is initial iff for every `Y` there is a unique morphism `X ⟶ Y`. -/ def isInitialEquivUnique (F : Discrete.{0} PEmpty.{1} ⥤ C) (X : C) : - IsColimit (⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ : Cocone F) ≃ ∀ Y : C, Unique (X ⟶ Y) where + IsColimit (⟨X, ⟨by aesop_cat, by simp⟩⟩ : Cocone F) ≃ ∀ Y : C, Unique (X ⟶ Y) where toFun t X := - { default := t.desc ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ - uniq := fun f => t.uniq ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ f (by aesop_cat) } + { default := t.desc ⟨X, ⟨by aesop_cat, by simp⟩⟩ + uniq := fun f => t.uniq ⟨X, ⟨by aesop_cat, by simp⟩⟩ f (by simp) } invFun u := { desc := fun s => (u s.pt).default uniq := fun s _ _ => (u s.pt).2 _ } @@ -152,7 +152,7 @@ def IsTerminal.from {X : C} (t : IsTerminal X) (Y : C) : Y ⟶ X := /-- Any two morphisms to a terminal object are equal. -/ theorem IsTerminal.hom_ext {X Y : C} (t : IsTerminal X) (f g : Y ⟶ X) : f = g := - IsLimit.hom_ext t (by aesop_cat) + IsLimit.hom_ext t (by simp) @[simp] theorem IsTerminal.comp_from {Z : C} (t : IsTerminal Z) {X Y : C} (f : X ⟶ Y) : @@ -169,7 +169,7 @@ def IsInitial.to {X : C} (t : IsInitial X) (Y : C) : X ⟶ Y := /-- Any two morphisms from an initial object are equal. -/ theorem IsInitial.hom_ext {X Y : C} (t : IsInitial X) (f g : X ⟶ Y) : f = g := - IsColimit.hom_ext t (by aesop_cat) + IsColimit.hom_ext t (by simp) @[simp] theorem IsInitial.to_comp {X : C} (t : IsInitial X) {Y Z : C} (f : Y ⟶ Z) : t.to Y ≫ f = t.to Z := @@ -217,12 +217,12 @@ variable (X : C) {F₁ : Discrete.{w} PEmpty ⥤ C} {F₂ : Discrete.{w'} PEmpty as long as the cone points are isomorphic. -/ def isLimitChangeEmptyCone {c₁ : Cone F₁} (hl : IsLimit c₁) (c₂ : Cone F₂) (hi : c₁.pt ≅ c₂.pt) : IsLimit c₂ where - lift c := hl.lift ⟨c.pt, by aesop_cat, by aesop_cat⟩ ≫ hi.hom + lift c := hl.lift ⟨c.pt, by aesop_cat, by simp⟩ ≫ hi.hom uniq c f _ := by dsimp rw [← hl.uniq _ (f ≫ hi.inv) _] · simp only [Category.assoc, Iso.inv_hom_id, Category.comp_id] - · aesop_cat + · simp /-- Replacing an empty cone in `IsLimit` by another with the same cone point is an equivalence. -/ @@ -239,12 +239,12 @@ def isLimitEmptyConeEquiv (c₁ : Cone F₁) (c₂ : Cone F₂) (h : c₁.pt ≅ as long as the cocone points are isomorphic. -/ def isColimitChangeEmptyCocone {c₁ : Cocone F₁} (hl : IsColimit c₁) (c₂ : Cocone F₂) (hi : c₁.pt ≅ c₂.pt) : IsColimit c₂ where - desc c := hi.inv ≫ hl.desc ⟨c.pt, by aesop_cat, by aesop_cat⟩ + desc c := hi.inv ≫ hl.desc ⟨c.pt, by aesop_cat, by simp⟩ uniq c f _ := by dsimp rw [← hl.uniq _ (hi.hom ≫ f) _] · simp only [Iso.inv_hom_id_assoc] - · aesop_cat + · simp /-- Replacing an empty cocone in `IsColimit` by another with the same cocone point is an equivalence. -/ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean b/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean index 202f266d2d66f..6e11689c3fc8b 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean @@ -275,7 +275,7 @@ theorem kernel.condition : kernel.ι f ≫ f = 0 := /-- The kernel built from `kernel.ι f` is limiting. -/ def kernelIsKernel : IsLimit (Fork.ofι (kernel.ι f) ((kernel.condition f).trans comp_zero.symm)) := - IsLimit.ofIsoLimit (limit.isLimit _) (Fork.ext (Iso.refl _) (by aesop_cat)) + IsLimit.ofIsoLimit (limit.isLimit _) (Fork.ext (Iso.refl _) (by simp)) /-- Given any morphism `k : W ⟶ X` satisfying `k ≫ f = 0`, `k` factors through `kernel.ι f` via `kernel.lift : W ⟶ kernel f`. -/ @@ -425,8 +425,8 @@ instance hasKernel_iso_comp {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso f] [H HasKernel (f ≫ g) where exists_limit := ⟨{ cone := KernelFork.ofι (kernel.ι g ≫ inv f) (by simp) - isLimit := isLimitAux _ (fun s => kernel.lift _ (s.ι ≫ f) (by aesop_cat)) - (by aesop_cat) fun s m w => by + isLimit := isLimitAux _ (fun s => kernel.lift _ (s.ι ≫ f) (by simp)) + (by simp) fun s m w => by simp_rw [← w] symm apply equalizer.hom_ext @@ -751,7 +751,7 @@ lemma colimit_ι_zero_cokernel_desc {C : Type*} [Category C] [HasZeroMorphisms C] {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) (h : f ≫ g = 0) [HasCokernel f] : colimit.ι (parallelPair f 0) WalkingParallelPair.zero ≫ cokernel.desc f g h = 0 := by rw [(colimit.w (parallelPair f 0) WalkingParallelPairHom.left).symm] - aesop_cat + simp @[simp] theorem cokernel.desc_zero {W : C} {h} : cokernel.desc f (0 : Y ⟶ W) h = 0 := by @@ -884,7 +884,7 @@ instance hasCokernel_comp_iso {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) [HasCokern isColimitAux _ (fun s => cokernel.desc _ (g ≫ s.π) (by rw [← Category.assoc, CokernelCofork.condition])) - (by aesop_cat) fun s m w => by + (by simp) fun s m w => by simp_rw [← w] symm apply coequalizer.hom_ext diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Products.lean b/Mathlib/CategoryTheory/Limits/Shapes/Products.lean index 02cc099007693..9cd40dec3a39b 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Products.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Products.lean @@ -565,7 +565,7 @@ instance {ι : Type*} (f : ι → Type*) (g : (i : ι) → (f i) → C) exists_limit := Nonempty.intro { cone := Fan.mk (∏ᶜ fun i => ∏ᶜ g i) (fun X => Pi.π (fun i => ∏ᶜ g i) X.1 ≫ Pi.π (g X.1) X.2) isLimit := mkFanLimit _ (fun s => Pi.lift fun b => Pi.lift fun c => s.proj ⟨b, c⟩) - (by aesop_cat) + (by simp) (by intro s m w; simp only [Fan.mk_pt]; symm; ext i x; simp_all [Sigma.forall]) } /-- An iterated product is a product over a sigma type. -/ @@ -586,7 +586,7 @@ instance {ι : Type*} (f : ι → Type*) (g : (i : ι) → (f i) → C) (fun X => Sigma.ι (g X.1) X.2 ≫ Sigma.ι (fun i => ∐ g i) X.1) isColimit := mkCofanColimit _ (fun s => Sigma.desc fun b => Sigma.desc fun c => s.inj ⟨b, c⟩) - (by aesop_cat) + (by simp) (by intro s m w; simp only [Cofan.mk_pt]; symm; ext i x; simp_all [Sigma.forall]) } /-- An iterated coproduct is a coproduct over a sigma type. -/ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean index 14f78f89b5b08..bc2c51f174f0e 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean @@ -89,27 +89,27 @@ theorem cocone_inr (s : CommSq f g h i) : s.cocone.inr = i := a commutative square identifies to the cocone of the flipped commutative square in the opposite category -/ def coneOp (p : CommSq f g h i) : p.cone.op ≅ p.flip.op.cocone := - PushoutCocone.ext (Iso.refl _) (by aesop_cat) (by aesop_cat) + PushoutCocone.ext (Iso.refl _) (by simp) (by simp) /-- The pullback cone in the opposite category associated to the cocone of a commutative square identifies to the cone of the flipped commutative square in the opposite category -/ def coconeOp (p : CommSq f g h i) : p.cocone.op ≅ p.flip.op.cone := - PullbackCone.ext (Iso.refl _) (by aesop_cat) (by aesop_cat) + PullbackCone.ext (Iso.refl _) (by simp) (by simp) /-- The pushout cocone obtained from the pullback cone associated to a commutative square in the opposite category identifies to the cocone associated to the flipped square. -/ def coneUnop {W X Y Z : Cᵒᵖ} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} (p : CommSq f g h i) : p.cone.unop ≅ p.flip.unop.cocone := - PushoutCocone.ext (Iso.refl _) (by aesop_cat) (by aesop_cat) + PushoutCocone.ext (Iso.refl _) (by simp) (by simp) /-- The pullback cone obtained from the pushout cone associated to a commutative square in the opposite category identifies to the cone associated to the flipped square. -/ def coconeUnop {W X Y Z : Cᵒᵖ} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} (p : CommSq f g h i) : p.cocone.unop ≅ p.flip.unop.cone := - PullbackCone.ext (Iso.refl _) (by aesop_cat) (by aesop_cat) + PullbackCone.ext (Iso.refl _) (by simp) (by simp) end CommSq @@ -221,7 +221,7 @@ lemma hom_ext (hP : IsPullback fst snd f g) {W : C} {k l : W ⟶ P} theorem of_isLimit {c : PullbackCone f g} (h : Limits.IsLimit c) : IsPullback c.fst c.snd f g := { w := c.condition isLimit' := ⟨IsLimit.ofIsoLimit h (Limits.PullbackCone.ext (Iso.refl _) - (by aesop_cat) (by aesop_cat))⟩ } + (by simp) (by simp))⟩ } /-- A variant of `of_isLimit` that is more useful with `apply`. -/ theorem of_isLimit' (w : CommSq fst snd f g) (h : Limits.IsLimit w.cone) : @@ -345,7 +345,7 @@ theorem of_horiz_isIso [IsIso fst] [IsIso g] (sq : CommSq fst snd f g) : IsPullb of_isLimit' sq (by refine - PullbackCone.IsLimit.mk _ (fun s => s.fst ≫ inv fst) (by aesop_cat) + PullbackCone.IsLimit.mk _ (fun s => s.fst ≫ inv fst) (by simp) (fun s => ?_) (by aesop_cat) simp only [← cancel_mono g, Category.assoc, ← sq.w, IsIso.inv_hom_id_assoc, s.condition]) @@ -431,7 +431,7 @@ theorem of_isColimit {c : PushoutCocone f g} (h : Limits.IsColimit c) : IsPushou { w := c.condition isColimit' := ⟨IsColimit.ofIsoColimit h (Limits.PushoutCocone.ext (Iso.refl _) - (by aesop_cat) (by aesop_cat))⟩ } + (by simp) (by simp))⟩ } /-- A variant of `of_isColimit` that is more useful with `apply`. -/ theorem of_isColimit' (w : CommSq f g inl inr) (h : Limits.IsColimit w.cocone) : @@ -1161,7 +1161,7 @@ theorem of_horiz_isIso [IsIso f] [IsIso inr] (sq : CommSq f g inl inr) : IsPusho (by refine PushoutCocone.IsColimit.mk _ (fun s => inv inr ≫ s.inr) (fun s => ?_) - (by aesop_cat) (by aesop_cat) + (by simp) (by simp) simp only [← cancel_epi f, s.condition, sq.w_assoc, IsIso.hom_inv_id_assoc]) theorem of_vert_isIso [IsIso g] [IsIso inl] (sq : CommSq f g inl inr) : IsPushout f g inl inr := diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Square.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Square.lean index 01bb011074096..ddd0fc6d43918 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Square.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Square.lean @@ -79,7 +79,7 @@ lemma IsPullback.of_iso {sq₁ sq₂ : Square C} (h : sq₁.IsPullback) refine CategoryTheory.IsPullback.of_iso h (evaluation₁.mapIso e) (evaluation₂.mapIso e) (evaluation₃.mapIso e) (evaluation₄.mapIso e) ?_ ?_ ?_ ?_ - all_goals aesop_cat + all_goals simp lemma IsPullback.iff_of_iso {sq₁ sq₂ : Square C} (e : sq₁ ≅ sq₂) : sq₁.IsPullback ↔ sq₂.IsPullback := @@ -90,7 +90,7 @@ lemma IsPushout.of_iso {sq₁ sq₂ : Square C} (h : sq₁.IsPushout) refine CategoryTheory.IsPushout.of_iso h (evaluation₁.mapIso e) (evaluation₂.mapIso e) (evaluation₃.mapIso e) (evaluation₄.mapIso e) ?_ ?_ ?_ ?_ - all_goals aesop_cat + all_goals simp lemma IsPushout.iff_of_iso {sq₁ sq₂ : Square C} (e : sq₁ ≅ sq₂) : sq₁.IsPushout ↔ sq₂.IsPushout := diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean b/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean index abcc9d116cda0..d1c5d29b44ebf 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean @@ -543,7 +543,7 @@ lemma reflexiveCoforkEquivCofork_functor_obj_π (G : ReflexiveCofork F) : ((reflexiveCoforkEquivCofork F).functor.obj G).π = G.π := by dsimp [reflexiveCoforkEquivCofork] rw [ReflexiveCofork.π, Cofork.π] - aesop_cat + simp @[simp] lemma reflexiveCoforkEquivCofork_inverse_obj_π diff --git a/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean b/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean index 2abdbdfda2ec4..e5f5d8bc1363b 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean @@ -164,11 +164,11 @@ end /-- A strong epimorphism that is a monomorphism is an isomorphism. -/ theorem isIso_of_mono_of_strongEpi (f : P ⟶ Q) [Mono f] [StrongEpi f] : IsIso f := - ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by aesop_cat⟩⟩ + ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by simp⟩⟩ /-- A strong monomorphism that is an epimorphism is an isomorphism. -/ theorem isIso_of_epi_of_strongMono (f : P ⟶ Q) [Epi f] [StrongMono f] : IsIso f := - ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by aesop_cat⟩⟩ + ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by simp⟩⟩ section diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean b/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean index d24c05e73cd2c..e3f907ea79b4a 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean @@ -43,7 +43,7 @@ section Univ variable (X : C) {F₁ : Discrete.{w} PEmpty ⥤ C} {F₂ : Discrete.{w'} PEmpty ⥤ C} theorem hasTerminalChangeDiagram (h : HasLimit F₁) : HasLimit F₂ := - ⟨⟨⟨⟨limit F₁, by aesop_cat, by aesop_cat⟩, + ⟨⟨⟨⟨limit F₁, by aesop_cat, by simp⟩, isLimitChangeEmptyCone C (limit.isLimit F₁) _ (eqToIso rfl)⟩⟩⟩ theorem hasTerminalChangeUniverse [h : HasLimitsOfShape (Discrete.{w} PEmpty) C] : @@ -51,7 +51,7 @@ theorem hasTerminalChangeUniverse [h : HasLimitsOfShape (Discrete.{w} PEmpty) C] has_limit _ := hasTerminalChangeDiagram C (h.1 (Functor.empty C)) theorem hasInitialChangeDiagram (h : HasColimit F₁) : HasColimit F₂ := - ⟨⟨⟨⟨colimit F₁, by aesop_cat, by aesop_cat⟩, + ⟨⟨⟨⟨colimit F₁, by aesop_cat, by simp⟩, isColimitChangeEmptyCocone C (colimit.isColimit F₁) _ (eqToIso rfl)⟩⟩⟩ theorem hasInitialChangeUniverse [h : HasColimitsOfShape (Discrete.{w} PEmpty) C] : @@ -92,7 +92,7 @@ theorem hasTerminal_of_unique (X : C) [∀ Y, Nonempty (Y ⟶ X)] [∀ Y, Subsin ⟨Classical.inhabited_of_nonempty', (Subsingleton.elim · _)⟩⟩ theorem IsTerminal.hasTerminal {X : C} (h : IsTerminal X) : HasTerminal C := - { has_limit := fun F => HasLimit.mk ⟨⟨X, by aesop_cat, by aesop_cat⟩, + { has_limit := fun F => HasLimit.mk ⟨⟨X, by aesop_cat, by simp⟩, isLimitChangeEmptyCone _ h _ (Iso.refl _)⟩ } /-- We can more explicitly show that a category has an initial object by specifying the object, @@ -104,7 +104,7 @@ theorem hasInitial_of_unique (X : C) [∀ Y, Nonempty (X ⟶ Y)] [∀ Y, Subsing theorem IsInitial.hasInitial {X : C} (h : IsInitial X) : HasInitial C where has_colimit F := - HasColimit.mk ⟨⟨X, by aesop_cat, by aesop_cat⟩, isColimitChangeEmptyCocone _ h _ (Iso.refl _)⟩ + HasColimit.mk ⟨⟨X, by aesop_cat, by simp⟩, isColimitChangeEmptyCocone _ h _ (Iso.refl _)⟩ /-- The map from an object to the terminal object. -/ abbrev terminal.from [HasTerminal C] (P : C) : P ⟶ ⊤_ C := diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean index f6dfbe28fd6cb..df7ea12153d1c 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean @@ -924,7 +924,7 @@ lemma pushoutCocone_inl_eq_inr_iff_of_isColimit {c : PushoutCocone f g} (hc : Is c.inl x₁ = c.inr x₂ ↔ ∃ (s : S), f s = x₁ ∧ g s = x₂ := by rw [pushoutCocone_inl_eq_inr_iff_of_iso (Cocones.ext (IsColimit.coconePointUniqueUpToIso hc (Pushout.isColimitCocone f g)) - (by aesop_cat))] + (by simp))] have := (mono_iff_injective f).2 h₁ apply Pushout.inl_eq_inr_iff diff --git a/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean b/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean index 4670206841758..06d17c50c7d91 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean @@ -330,9 +330,9 @@ theorem eq_lift_of_comp_eq (g : X ⟶ widePullback _ _ arrows) : · apply h1 theorem hom_eq_lift (g : X ⟶ widePullback _ _ arrows) : - g = lift (g ≫ base arrows) (fun j => g ≫ π arrows j) (by aesop_cat) := by + g = lift (g ≫ base arrows) (fun j => g ≫ π arrows j) (by simp) := by apply eq_lift_of_comp_eq - · aesop_cat + · simp · rfl -- Porting note: quite a few missing refl's in aesop_cat now @[ext 1100] @@ -400,7 +400,7 @@ theorem hom_eq_desc (g : widePushout _ _ arrows ⟶ X) : rw [← Category.assoc] simp := by apply eq_desc_of_comp_eq - · aesop_cat + · simp · rfl -- Porting note: another missing rfl @[ext 1100] diff --git a/Mathlib/CategoryTheory/Localization/Bifunctor.lean b/Mathlib/CategoryTheory/Localization/Bifunctor.lean index e4556b6e608d8..7755ae9cca22e 100644 --- a/Mathlib/CategoryTheory/Localization/Bifunctor.lean +++ b/Mathlib/CategoryTheory/Localization/Bifunctor.lean @@ -168,8 +168,8 @@ and `Lifting₂ L₁ L₂ W₁ W₂ F₂ F₂'` hold. -/ noncomputable def lift₂NatIso (e : F₁ ≅ F₂) : F₁' ≅ F₂' where hom := lift₂NatTrans L₁ L₂ W₁ W₂ F₁ F₂ F₁' F₂' e.hom inv := lift₂NatTrans L₁ L₂ W₁ W₂ F₂ F₁ F₂' F₁' e.inv - hom_inv_id := natTrans₂_ext L₁ L₂ W₁ W₂ (by aesop_cat) - inv_hom_id := natTrans₂_ext L₁ L₂ W₁ W₂ (by aesop_cat) + hom_inv_id := natTrans₂_ext L₁ L₂ W₁ W₂ (by simp) + inv_hom_id := natTrans₂_ext L₁ L₂ W₁ W₂ (by simp) end diff --git a/Mathlib/CategoryTheory/Localization/Construction.lean b/Mathlib/CategoryTheory/Localization/Construction.lean index b82770a4d5037..5bcd82a33e3e7 100644 --- a/Mathlib/CategoryTheory/Localization/Construction.lean +++ b/Mathlib/CategoryTheory/Localization/Construction.lean @@ -154,7 +154,7 @@ def lift : W.Localization ⥤ D := -- Porting note: rest of proof was `rcases r with ⟨⟩; tidy` rcases r with (_|_|⟨f,hf⟩|⟨f,hf⟩) · aesop_cat - · aesop_cat + · simp all_goals dsimp haveI := hG f hf diff --git a/Mathlib/CategoryTheory/Localization/DerivabilityStructure/Constructor.lean b/Mathlib/CategoryTheory/Localization/DerivabilityStructure/Constructor.lean index 2adefbca8f183..b7da168c46084 100644 --- a/Mathlib/CategoryTheory/Localization/DerivabilityStructure/Constructor.lean +++ b/Mathlib/CategoryTheory/Localization/DerivabilityStructure/Constructor.lean @@ -90,7 +90,7 @@ lemma isConnected : (isPreconnected_zigzag (RightResolution.mk (𝟙 _) (W₂.id_mem _)) (RightResolution.mk ρ.w.right ρ.hw.2)) refine Zigzag.trans ?_ (Zigzag.trans this ?_) - · exact Zigzag.of_hom (eqToHom (by aesop)) + · exact Zigzag.of_hom (eqToHom (by simp)) · apply Zigzag.of_inv refine CostructuredArrow.homMk (StructuredArrow.homMk ρ.X₁.hom (by simp)) ?_ ext diff --git a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean index b2025fc3fd6e4..2b1164339e414 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean @@ -236,7 +236,7 @@ noncomputable def braidedCategoryOfFullyFaithful {C D : Type*} [Category C] [Cat braidedCategoryOfFaithful F (fun X Y => F.preimageIso ((μIso F _ _).symm ≪≫ β_ (F.obj X) (F.obj Y) ≪≫ (μIso F _ _))) - (by aesop_cat) + (by simp) section diff --git a/Mathlib/CategoryTheory/Monoidal/End.lean b/Mathlib/CategoryTheory/Monoidal/End.lean index ab0ce6bcf4ec1..2994d7b5dc0ef 100644 --- a/Mathlib/CategoryTheory/Monoidal/End.lean +++ b/Mathlib/CategoryTheory/Monoidal/End.lean @@ -141,7 +141,7 @@ theorem ε_naturality {X Y : C} (f : X ⟶ Y) [F.LaxMonoidal] : @[reassoc (attr := simp)] theorem η_naturality {X Y : C} (f : X ⟶ Y) [F.OplaxMonoidal]: (η F).app X ≫ (𝟙_ (C ⥤ C)).map f = (η F).app X ≫ f := by - aesop_cat + simp @[reassoc (attr := simp)] theorem μ_naturality {m n : M} {X Y : C} (f : X ⟶ Y) [F.LaxMonoidal] : diff --git a/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean b/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean index 84cdce3c91ae3..5125b23e93c80 100644 --- a/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean +++ b/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean @@ -52,7 +52,7 @@ def limitCone (F : J ⥤ Mon_ C) : Cone F where -/ def forgetMapConeLimitConeIso (F : J ⥤ Mon_ C) : (forget C).mapCone (limitCone F) ≅ limit.cone (F ⋙ forget C) := - Cones.ext (Iso.refl _) (by aesop_cat) + Cones.ext (Iso.refl _) (by simp) /-- Implementation of `Mon_.hasLimitsOfShape`: the proposed cone over a functor `F : J ⥤ Mon_ C` is a limit cone. diff --git a/Mathlib/CategoryTheory/Monoidal/Subcategory.lean b/Mathlib/CategoryTheory/Monoidal/Subcategory.lean index 52293104297ba..407983573fdc0 100644 --- a/Mathlib/CategoryTheory/Monoidal/Subcategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/Subcategory.lean @@ -128,7 +128,7 @@ instance fullBraidedSubcategory : BraidedCategory (FullSubcategory P) := braidedCategoryOfFaithful (fullSubcategoryInclusion P) (fun X Y => ⟨(β_ X.1 Y.1).hom, (β_ X.1 Y.1).inv, (β_ X.1 Y.1).hom_inv_id, (β_ X.1 Y.1).inv_hom_id⟩) - fun X Y => by aesop_cat + fun X Y => by simp /-- The forgetful braided functor from a full braided subcategory into the original category ("forgetting" the condition). diff --git a/Mathlib/CategoryTheory/NatIso.lean b/Mathlib/CategoryTheory/NatIso.lean index e12d02fb99a67..caa4309d21b37 100644 --- a/Mathlib/CategoryTheory/NatIso.lean +++ b/Mathlib/CategoryTheory/NatIso.lean @@ -227,7 +227,7 @@ theorem ofComponents.app (app' : ∀ X : C, F.obj X ≅ G.obj X) (naturality) (X /-- A natural transformation is an isomorphism if all its components are isomorphisms. -/ theorem isIso_of_isIso_app (α : F ⟶ G) [∀ X : C, IsIso (α.app X)] : IsIso α := - (ofComponents (fun X => asIso (α.app X)) (by aesop)).isIso_hom + (ofComponents (fun X => asIso (α.app X)) (by simp)).isIso_hom /-- Horizontal composition of natural isomorphisms. -/ @[simps] diff --git a/Mathlib/CategoryTheory/Opposites.lean b/Mathlib/CategoryTheory/Opposites.lean index 8cb8e3e027660..914d995a8d4cd 100644 --- a/Mathlib/CategoryTheory/Opposites.lean +++ b/Mathlib/CategoryTheory/Opposites.lean @@ -120,7 +120,7 @@ end /-- If `f` is an isomorphism, so is `f.op` -/ instance isIso_op {X Y : C} (f : X ⟶ Y) [IsIso f] : IsIso f.op := - ⟨⟨(inv f).op, ⟨Quiver.Hom.unop_inj (by aesop_cat), Quiver.Hom.unop_inj (by aesop_cat)⟩⟩⟩ + ⟨⟨(inv f).op, ⟨Quiver.Hom.unop_inj (by simp), Quiver.Hom.unop_inj (by simp)⟩⟩⟩ /-- If `f.op` is an isomorphism `f` must be too. (This cannot be an instance as it would immediately loop!) diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean index 620eae9dd1af8..0b47b17e1e3a3 100644 --- a/Mathlib/CategoryTheory/Pi/Basic.lean +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -109,7 +109,7 @@ def comapComp (f : K → J) (g : J → I) : comap C g ⋙ comap (C ∘ g) f ≅ /-- The natural isomorphism between pulling back then evaluating, and just evaluating. -/ @[simps!] def comapEvalIsoEval (h : J → I) (j : J) : comap C h ⋙ eval (C ∘ h) j ≅ eval C (h j) := - NatIso.ofComponents (fun _ => Iso.refl _) (by simp only [Iso.refl]; aesop_cat) + NatIso.ofComponents (fun _ => Iso.refl _) (by simp only [Iso.refl]; simp) end diff --git a/Mathlib/CategoryTheory/Preadditive/Biproducts.lean b/Mathlib/CategoryTheory/Preadditive/Biproducts.lean index 49104d0c125a5..4a9c75efd68f9 100644 --- a/Mathlib/CategoryTheory/Preadditive/Biproducts.lean +++ b/Mathlib/CategoryTheory/Preadditive/Biproducts.lean @@ -143,7 +143,7 @@ def isBilimitOfIsLimit {f : J → C} (t : Bicone f) (ht : IsLimit t.toCone) : t. /-- We can turn any limit cone over a pair into a bilimit bicone. -/ def biconeIsBilimitOfLimitConeOfIsLimit {f : J → C} {t : Cone (Discrete.functor f)} (ht : IsLimit t) : (Bicone.ofLimitCone ht).IsBilimit := - isBilimitOfIsLimit _ <| IsLimit.ofIsoLimit ht <| Cones.ext (Iso.refl _) (by aesop_cat) + isBilimitOfIsLimit _ <| IsLimit.ofIsoLimit ht <| Cones.ext (Iso.refl _) (by simp) /-- In a preadditive category, any finite bicone which is a colimit cocone is in fact a bilimit bicone. -/ diff --git a/Mathlib/CategoryTheory/Preadditive/OfBiproducts.lean b/Mathlib/CategoryTheory/Preadditive/OfBiproducts.lean index c2f2f2b51b091..499eed7f4a959 100644 --- a/Mathlib/CategoryTheory/Preadditive/OfBiproducts.lean +++ b/Mathlib/CategoryTheory/Preadditive/OfBiproducts.lean @@ -51,12 +51,12 @@ theorem isUnital_leftAdd : EckmannHilton.IsUnital (· +ₗ ·) 0 := by have hr : ∀ f : X ⟶ Y, biprod.lift (0 : X ⟶ Y) f = f ≫ biprod.inr := by intro f ext - · aesop_cat + · simp · simp [biprod.lift_fst, Category.assoc, biprod.inr_fst, comp_zero] have hl : ∀ f : X ⟶ Y, biprod.lift f (0 : X ⟶ Y) = f ≫ biprod.inl := by intro f ext - · aesop_cat + · simp · simp [biprod.lift_snd, Category.assoc, biprod.inl_snd, comp_zero] exact { left_id := fun f => by simp [hr f, leftAdd, Category.assoc, Category.comp_id, biprod.inr_desc], @@ -67,12 +67,12 @@ theorem isUnital_rightAdd : EckmannHilton.IsUnital (· +ᵣ ·) 0 := by have h₂ : ∀ f : X ⟶ Y, biprod.desc (0 : X ⟶ Y) f = biprod.snd ≫ f := by intro f ext - · aesop_cat + · simp · simp only [biprod.inr_desc, BinaryBicone.inr_snd_assoc] have h₁ : ∀ f : X ⟶ Y, biprod.desc f (0 : X ⟶ Y) = biprod.fst ≫ f := by intro f ext - · aesop_cat + · simp · simp only [biprod.inr_desc, BinaryBicone.inr_fst_assoc, zero_comp] exact { left_id := fun f => by simp [h₂ f, rightAdd, biprod.lift_snd_assoc, Category.id_comp], diff --git a/Mathlib/CategoryTheory/Simple.lean b/Mathlib/CategoryTheory/Simple.lean index 619a1bdfe7847..3d1d618bb6c68 100644 --- a/Mathlib/CategoryTheory/Simple.lean +++ b/Mathlib/CategoryTheory/Simple.lean @@ -116,7 +116,7 @@ variable (C) /-- We don't want the definition of 'simple' to include the zero object, so we check that here. -/ theorem zero_not_simple [Simple (0 : C)] : False := - (Simple.mono_isIso_iff_nonzero (0 : (0 : C) ⟶ (0 : C))).mp ⟨⟨0, by aesop_cat⟩⟩ rfl + (Simple.mono_isIso_iff_nonzero (0 : (0 : C) ⟶ (0 : C))).mp ⟨⟨0, by simp⟩⟩ rfl end diff --git a/Mathlib/CategoryTheory/SingleObj.lean b/Mathlib/CategoryTheory/SingleObj.lean index a3e3ca9146610..48fed64ab7d7d 100644 --- a/Mathlib/CategoryTheory/SingleObj.lean +++ b/Mathlib/CategoryTheory/SingleObj.lean @@ -205,11 +205,11 @@ def toSingleObjEquiv (e : M ≃* N) : SingleObj M ≌ SingleObj N where unitIso := eqToIso (by rw [← MonoidHom.comp_toFunctor, ← MonoidHom.id_toFunctor] congr 1 - aesop_cat) + simp) counitIso := eqToIso (by rw [← MonoidHom.comp_toFunctor, ← MonoidHom.id_toFunctor] congr 1 - aesop_cat) + simp) end MulEquiv diff --git a/Mathlib/CategoryTheory/Sites/Pretopology.lean b/Mathlib/CategoryTheory/Sites/Pretopology.lean index d888a2611acf2..bde2d0fb7bd77 100644 --- a/Mathlib/CategoryTheory/Sites/Pretopology.lean +++ b/Mathlib/CategoryTheory/Sites/Pretopology.lean @@ -172,7 +172,7 @@ def trivial : Pretopology C where pullbacks X Y f S := by rintro ⟨Z, g, i, rfl⟩ refine ⟨pullback g f, pullback.snd _ _, ?_, ?_⟩ - · refine ⟨⟨pullback.lift (f ≫ inv g) (𝟙 _) (by simp), ⟨?_, by aesop_cat⟩⟩⟩ + · refine ⟨⟨pullback.lift (f ≫ inv g) (𝟙 _) (by simp), ⟨?_, by simp⟩⟩⟩ ext · rw [assoc, pullback.lift_fst, ← pullback.condition_assoc] simp diff --git a/Mathlib/CategoryTheory/Sites/Subsheaf.lean b/Mathlib/CategoryTheory/Sites/Subsheaf.lean index be30cc3fa8ebc..421d09f0798b1 100644 --- a/Mathlib/CategoryTheory/Sites/Subsheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Subsheaf.lean @@ -57,7 +57,7 @@ instance : PartialOrder (Subpresheaf F) := PartialOrder.lift Subpresheaf.obj (fun _ _ => Subpresheaf.ext) instance : Top (Subpresheaf F) := - ⟨⟨fun _ => ⊤, @fun U _ _ x _ => by aesop_cat⟩⟩ + ⟨⟨fun _ => ⊤, @fun U _ _ x _ => by simp⟩⟩ instance : Nonempty (Subpresheaf F) := inferInstance diff --git a/Mathlib/CategoryTheory/Subobject/Comma.lean b/Mathlib/CategoryTheory/Subobject/Comma.lean index 20c1832042687..31c43bb091ef1 100644 --- a/Mathlib/CategoryTheory/Subobject/Comma.lean +++ b/Mathlib/CategoryTheory/Subobject/Comma.lean @@ -107,7 +107,7 @@ def subobjectEquiv [HasFiniteLimits C] [PreservesFiniteLimits T] (A : Structured refine ⟨fun h => Subobject.mk_le_mk_of_comm ?_ ?_, fun h => ?_⟩ · exact homMk (Subobject.ofMkLEMk _ _ h) ((cancel_mono (T.map g.right)).1 (by simp [← T.map_comp])) - · aesop_cat + · simp · refine Subobject.mk_le_mk_of_comm (Subobject.ofMkLEMk _ _ h).right ?_ exact congr_arg CommaMorphism.right (Subobject.ofMkLEMk_comp h) @@ -187,7 +187,7 @@ theorem lift_projectQuotient [HasFiniteColimits C] [PreservesFiniteColimits S] · exact (Subobject.underlyingIso f.unop.left.op).unop · refine (cancel_epi (S.map f.unop.left)).1 ?_ simpa [← Category.assoc, ← S.map_comp] using hq - · exact Quiver.Hom.unop_inj (by aesop_cat)) + · exact Quiver.Hom.unop_inj (by simp)) /-- Technical lemma for `quotientEquiv`. -/ theorem unop_left_comp_ofMkLEMk_unop {A : CostructuredArrow S T} {P Q : (CostructuredArrow S T)ᵒᵖ} diff --git a/Mathlib/CategoryTheory/Subobject/Lattice.lean b/Mathlib/CategoryTheory/Subobject/Lattice.lean index 813f8887842b6..70d9b061a97da 100644 --- a/Mathlib/CategoryTheory/Subobject/Lattice.lean +++ b/Mathlib/CategoryTheory/Subobject/Lattice.lean @@ -59,7 +59,7 @@ variable [HasPullbacks C] is (isomorphic to) the top object in `MonoOver X`. -/ def pullbackTop (f : X ⟶ Y) : (pullback f).obj ⊤ ≅ ⊤ := iso_of_both_ways (leTop _) - (homMk (pullback.lift f (𝟙 _) (by aesop_cat)) (pullback.lift_snd _ _ _)) + (homMk (pullback.lift f (𝟙 _) (by simp)) (pullback.lift_snd _ _ _)) /-- There is a morphism from `⊤ : MonoOver A` to the pullback of a monomorphism along itself; as the category is thin this is an isomorphism. -/ @@ -527,7 +527,7 @@ def leInfCone {A : C} (s : Set (Subobject A)) (f : Subobject A) (k : ∀ g ∈ s (by rcases j with ⟨-, ⟨g, ⟨m, rfl⟩⟩⟩ simpa using m)))) - (by aesop_cat) + (by simp) @[simp] theorem leInfCone_π_app_none {A : C} (s : Set (Subobject A)) (f : Subobject A) @@ -572,7 +572,7 @@ theorem sInf_le {A : C} (s : Set (Subobject A)) (f) (hf : f ∈ s) : sInf s ≤ simp only [Category.comp_id, Category.assoc, ← underlyingIso_hom_comp_eq_mk, Subobject.arrow_congr, congrArg_mpr_hom_left, Iso.cancel_iso_hom_left] convert limit.w (wideCospan s) (WidePullbackShape.Hom.term _) - aesop_cat + simp theorem le_sInf {A : C} (s : Set (Subobject A)) (f : Subobject A) (k : ∀ g ∈ s, f ≤ g) : f ≤ sInf s := by diff --git a/Mathlib/CategoryTheory/Triangulated/Basic.lean b/Mathlib/CategoryTheory/Triangulated/Basic.lean index e76c664d5f1d3..b8e3b05949917 100644 --- a/Mathlib/CategoryTheory/Triangulated/Basic.lean +++ b/Mathlib/CategoryTheory/Triangulated/Basic.lean @@ -249,7 +249,7 @@ def binaryProductTriangleIsoBinaryBiproductTriangle (X₁ X₂ : C) [HasZeroMorphisms C] [HasBinaryBiproduct X₁ X₂] : binaryProductTriangle X₁ X₂ ≅ binaryBiproductTriangle X₁ X₂ := Triangle.isoMk _ _ (Iso.refl _) (biprod.isoProd X₁ X₂).symm (Iso.refl _) - (by aesop_cat) (by aesop_cat) (by aesop_cat) + (by aesop_cat) (by simp) (by simp) section diff --git a/Mathlib/CategoryTheory/Triangulated/Functor.lean b/Mathlib/CategoryTheory/Triangulated/Functor.lean index c7c860c35beec..faa14d9b0ea07 100644 --- a/Mathlib/CategoryTheory/Triangulated/Functor.lean +++ b/Mathlib/CategoryTheory/Triangulated/Functor.lean @@ -76,7 +76,7 @@ noncomputable def mapTriangleCommShiftIso (n : ℤ) : Triangle.shiftFunctor C n ⋙ F.mapTriangle ≅ F.mapTriangle ⋙ Triangle.shiftFunctor D n := NatIso.ofComponents (fun T => Triangle.isoMk _ _ ((F.commShiftIso n).app _) ((F.commShiftIso n).app _) ((F.commShiftIso n).app _) - (by aesop_cat) (by aesop_cat) (by + (by simp) (by simp) (by dsimp simp only [map_units_smul, map_comp, Linear.units_smul_comp, assoc, Linear.comp_units_smul, ← F.commShiftIso_hom_naturality_assoc] @@ -103,7 +103,7 @@ def mapTriangleRotateIso : NatIso.ofComponents (fun T => Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) ((F.commShiftIso (1 : ℤ)).symm.app _) - (by aesop_cat) (by aesop_cat) (by aesop_cat)) (by aesop_cat) + (by simp) (by simp) (by simp)) (by aesop_cat) /-- `F.mapTriangle` commutes with the inverse of the rotation of triangles. -/ @[simps!] @@ -112,7 +112,7 @@ noncomputable def mapTriangleInvRotateIso [F.Additive] : Pretriangulated.invRotate C ⋙ F.mapTriangle := NatIso.ofComponents (fun T => Triangle.isoMk _ _ ((F.commShiftIso (-1 : ℤ)).symm.app _) (Iso.refl _) (Iso.refl _) - (by aesop_cat) (by aesop_cat) (by aesop_cat)) (by aesop_cat) + (by simp) (by simp) (by simp)) (by aesop_cat) variable (C) in diff --git a/Mathlib/CategoryTheory/Triangulated/Opposite/Pretriangulated.lean b/Mathlib/CategoryTheory/Triangulated/Opposite/Pretriangulated.lean index 5e146b7317e28..dcec2c68e78ac 100644 --- a/Mathlib/CategoryTheory/Triangulated/Opposite/Pretriangulated.lean +++ b/Mathlib/CategoryTheory/Triangulated/Opposite/Pretriangulated.lean @@ -88,7 +88,7 @@ noncomputable def contractibleTriangleIso (X : Cᵒᵖ) : rw [IsZero.iff_id_eq_zero] change (𝟙 ((0 : C)⟦(-1 : ℤ)⟧)).op = 0 rw [← Functor.map_id, id_zero, Functor.map_zero, op_zero])) - (by aesop_cat) (by aesop_cat) (by aesop_cat) + (by simp) (by simp) (by simp) lemma contractible_distinguished (X : Cᵒᵖ) : contractibleTriangle X ∈ distinguishedTriangles C := by @@ -103,7 +103,7 @@ noncomputable def rotateTriangleOpEquivalenceInverseObjRotateUnopIso (T : Triang ((triangleOpEquivalence C).inverse.obj T).unop := Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (-((opShiftFunctorEquivalence C 1).unitIso.app T.obj₁).unop) (by simp) - (Quiver.Hom.op_inj (by aesop_cat)) (by aesop_cat) + (Quiver.Hom.op_inj (by simp)) (by simp) lemma rotate_distinguished_triangle (T : Triangle Cᵒᵖ) : T ∈ distinguishedTriangles C ↔ T.rotate ∈ distinguishedTriangles C := by @@ -119,7 +119,7 @@ lemma distinguished_cocone_triangle {X Y : Cᵒᵖ} (f : X ⟶ Y) : (shiftFunctor Cᵒᵖ (1 : ℤ)).map h.op, ?_⟩ simp only [mem_distinguishedTriangles_iff] refine Pretriangulated.isomorphic_distinguished _ H _ ?_ - exact Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) (by aesop_cat) (by aesop_cat) + exact Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) (by simp) (by simp) (Quiver.Hom.op_inj (by simp [shift_unop_opShiftFunctorEquivalence_counitIso_inv_app])) lemma complete_distinguished_triangle_morphism (T₁ T₂ : Triangle Cᵒᵖ) diff --git a/Mathlib/CategoryTheory/Triangulated/Opposite/Triangle.lean b/Mathlib/CategoryTheory/Triangulated/Opposite/Triangle.lean index cd0a07b4be3dd..2618748d4f279 100644 --- a/Mathlib/CategoryTheory/Triangulated/Opposite/Triangle.lean +++ b/Mathlib/CategoryTheory/Triangulated/Opposite/Triangle.lean @@ -73,7 +73,7 @@ equivalence `triangleOpEquivalence C : (Triangle C)ᵒᵖ ≌ Triangle Cᵒᵖ` @[simps!] noncomputable def unitIso : 𝟭 _ ≅ functor C ⋙ inverse C := NatIso.ofComponents (fun T => Iso.op - (Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) (by aesop_cat) (by aesop_cat) + (Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) (by simp) (by simp) (Quiver.Hom.op_inj (by simp [shift_unop_opShiftFunctorEquivalence_counitIso_inv_app])))) (fun {T₁ T₂} f => Quiver.Hom.unop_inj (by aesop_cat)) @@ -84,8 +84,8 @@ equivalence `triangleOpEquivalence C : (Triangle C)ᵒᵖ ≌ Triangle Cᵒᵖ` noncomputable def counitIso : inverse C ⋙ functor C ≅ 𝟭 _ := NatIso.ofComponents (fun T => by refine Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) ?_ ?_ ?_ - · aesop_cat - · aesop_cat + · simp + · simp · dsimp rw [Functor.map_id, comp_id, id_comp, Functor.map_comp, ← opShiftFunctorEquivalence_counitIso_inv_naturality_assoc, diff --git a/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean b/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean index c3b954b379632..1cbad40237ea8 100644 --- a/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean +++ b/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean @@ -219,7 +219,7 @@ lemma contractible_distinguished₁ (X : C) : refine isomorphic_distinguished _ (inv_rot_of_distTriang _ (contractible_distinguished X)) _ ?_ exact Triangle.isoMk _ _ (Functor.mapZeroObject _).symm (Iso.refl _) (Iso.refl _) - (by aesop_cat) (by aesop_cat) (by aesop_cat) + (by simp) (by simp) (by simp) /-- Obvious triangles `X ⟶ 0 ⟶ X⟦1⟧ ⟶ X⟦1⟧` are distinguished -/ lemma contractible_distinguished₂ (X : C) : @@ -227,7 +227,7 @@ lemma contractible_distinguished₂ (X : C) : refine isomorphic_distinguished _ (inv_rot_of_distTriang _ (contractible_distinguished₁ (X⟦(1 : ℤ)⟧))) _ ?_ exact Triangle.isoMk _ _ ((shiftEquiv C (1 : ℤ)).unitIso.app X) (Iso.refl _) (Iso.refl _) - (by aesop_cat) (by aesop_cat) + (by simp) (by simp) (by dsimp; simp only [shift_shiftFunctorCompIsoId_inv_app, id_comp]) namespace Triangle @@ -531,7 +531,7 @@ lemma binaryBiproductTriangle_distinguished (X₁ X₂ : C) : dsimp at he₁ he₂ refine isomorphic_distinguished _ mem _ (Iso.symm ?_) refine Triangle.isoMk _ _ (Iso.refl _) e (Iso.refl _) - (by aesop_cat) (by aesop_cat) (by aesop_cat) + (by aesop_cat) (by aesop_cat) (by simp) lemma binaryProductTriangle_distinguished (X₁ X₂ : C) : binaryProductTriangle X₁ X₂ ∈ distTriang C := diff --git a/Mathlib/CategoryTheory/Triangulated/Subcategory.lean b/Mathlib/CategoryTheory/Triangulated/Subcategory.lean index 22baf9e9ef9b8..f662bd84a092a 100644 --- a/Mathlib/CategoryTheory/Triangulated/Subcategory.lean +++ b/Mathlib/CategoryTheory/Triangulated/Subcategory.lean @@ -81,7 +81,7 @@ def isoClosure : Subcategory C where exact le_isoClosure _ _ (S.ext₂' (Triangle.mk (e₁.inv ≫ T.mor₁) (T.mor₂ ≫ e₃.hom) (e₃.inv ≫ T.mor₃ ≫ e₁.hom⟦1⟧')) (isomorphic_distinguished _ hT _ - (Triangle.isoMk _ _ e₁.symm (Iso.refl _) e₃.symm (by aesop_cat) (by aesop_cat) (by + (Triangle.isoMk _ _ e₁.symm (Iso.refl _) e₃.symm (by simp) (by simp) (by dsimp simp only [assoc, Iso.cancel_iso_inv_left, ← Functor.map_comp, e₁.hom_inv_id, Functor.map_id, comp_id]))) h₁ h₃) @@ -155,7 +155,7 @@ instance respectsIso_W : S.W.RespectsIso where precomp {X' X Y} e (he : IsIso e) := by rintro f ⟨Z, g, h, mem, mem'⟩ refine ⟨Z, g, h ≫ inv e⟦(1 : ℤ)⟧', isomorphic_distinguished _ mem _ ?_, mem'⟩ - refine Triangle.isoMk _ _ (asIso e) (Iso.refl _) (Iso.refl _) (by aesop_cat) (by aesop_cat) ?_ + refine Triangle.isoMk _ _ (asIso e) (Iso.refl _) (Iso.refl _) (by simp) (by simp) ?_ dsimp simp only [Functor.map_inv, assoc, IsIso.inv_hom_id, comp_id, id_comp] postcomp {X Y Y'} e (he : IsIso e) := by diff --git a/Mathlib/CategoryTheory/Triangulated/TriangleShift.lean b/Mathlib/CategoryTheory/Triangulated/TriangleShift.lean index b811588238be2..c06cd41d963ae 100644 --- a/Mathlib/CategoryTheory/Triangulated/TriangleShift.lean +++ b/Mathlib/CategoryTheory/Triangulated/TriangleShift.lean @@ -65,7 +65,7 @@ noncomputable def Triangle.shiftFunctorZero : Triangle.shiftFunctor C 0 ≅ 𝟭 NatIso.ofComponents (fun T => Triangle.isoMk _ _ ((CategoryTheory.shiftFunctorZero C ℤ).app _) ((CategoryTheory.shiftFunctorZero C ℤ).app _) ((CategoryTheory.shiftFunctorZero C ℤ).app _) - (by aesop_cat) (by aesop_cat) (by + (by simp) (by simp) (by dsimp simp only [one_smul, assoc, shiftFunctorComm_zero_hom_app, ← Functor.map_comp, Iso.inv_hom_id_app, Functor.id_obj, Functor.map_id, @@ -109,7 +109,7 @@ noncomputable def rotateRotateRotateIso : rotate C ⋙ rotate C ⋙ rotate C ≅ Triangle.shiftFunctor C 1 := NatIso.ofComponents (fun T => Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) - (by aesop_cat) (by aesop_cat) (by aesop_cat)) + (by simp) (by simp) (by simp)) (by aesop_cat) /-- Rotating triangles three times backwards identifies with the shift by `-1`. -/ @@ -117,8 +117,8 @@ noncomputable def invRotateInvRotateInvRotateIso : invRotate C ⋙ invRotate C ⋙ invRotate C ≅ Triangle.shiftFunctor C (-1) := NatIso.ofComponents (fun T => Triangle.isoMk _ _ (Iso.refl _) (Iso.refl _) (Iso.refl _) - (by aesop_cat) - (by aesop_cat) + (by simp) + (by simp) (by dsimp [shiftFunctorCompIsoId] simp [shiftFunctorComm_eq C _ _ _ (add_neg_cancel (1 : ℤ))])) diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index 05d0c28efabb3..e362440ade3cf 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -523,7 +523,7 @@ def yonedaCompUliftFunctorEquiv (F : Cᵒᵖ ⥤ Type max v₁ w) (X : C) : dsimp rw [Category.comp_id] rfl - right_inv f := by aesop_cat + right_inv f := by simp /-- The Yoneda lemma asserts that the Yoneda pairing `(X : Cᵒᵖ, F : Cᵒᵖ ⥤ Type) ↦ (yoneda.obj (unop X) ⟶ F)` @@ -705,7 +705,7 @@ def coyonedaCompUliftFunctorEquiv (F : C ⥤ Type max v₁ w) (X : Cᵒᵖ) : dsimp rw [Category.id_comp] rfl - right_inv f := by aesop_cat + right_inv f := by simp /-- The Coyoneda lemma asserts that the Coyoneda pairing `(X : C, F : C ⥤ Type) ↦ (coyoneda.obj X ⟶ F)` diff --git a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean index 7cd217fadb8cf..0031c9b006e43 100644 --- a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean +++ b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean @@ -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 simp)) 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 simp)) 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/Data/MLList/BestFirst.lean b/Mathlib/Data/MLList/BestFirst.lean index 430ac91fa3323..4f0ae77aca908 100644 --- a/Mathlib/Data/MLList/BestFirst.lean +++ b/Mathlib/Data/MLList/BestFirst.lean @@ -307,7 +307,7 @@ variable {m : Type → Type} {α : Type} [Monad m] [Alternative m] [LinearOrder for simple use cases. -/ local instance instOrderBotEq {a : α} : OrderBot { x : α // x = a } where bot := ⟨a, rfl⟩ - bot_le := by aesop + bot_le := by simp /-- A lazy list recording the best first search of a graph generated by a function diff --git a/Mathlib/Data/Matrix/ColumnRowPartitioned.lean b/Mathlib/Data/Matrix/ColumnRowPartitioned.lean index c583ebb1d3b5e..d4d5448a54f8f 100644 --- a/Mathlib/Data/Matrix/ColumnRowPartitioned.lean +++ b/Mathlib/Data/Matrix/ColumnRowPartitioned.lean @@ -125,8 +125,7 @@ lemma fromRows_toRows (A : Matrix (m₁ ⊕ m₂) n R) : fromRows A.toRows₁ A. lemma fromRows_inj : Function.Injective2 (@fromRows R m₁ m₂ n) := by intros x1 x2 y1 y2 - simp only [funext_iff, ← Matrix.ext_iff] - aesop + simp [← Matrix.ext_iff] lemma fromCols_inj : Function.Injective2 (@fromCols R m n₁ n₂) := by intros x1 x2 y1 y2 diff --git a/Mathlib/Data/Matroid/Closure.lean b/Mathlib/Data/Matroid/Closure.lean index cb14ccdda92a3..534a3d8ddbc12 100644 --- a/Mathlib/Data/Matroid/Closure.lean +++ b/Mathlib/Data/Matroid/Closure.lean @@ -154,7 +154,7 @@ lemma closure_subset_ground (M : Matroid α) (X : Set α) : M.closure X ⊆ M.E simp_rw [closure_def, inter_assoc, inter_self] lemma inter_ground_subset_closure (M : Matroid α) (X : Set α) : X ∩ M.E ⊆ M.closure X := by - simp_rw [closure_def, subset_sInter_iff]; aesop + simp_rw [closure_def, subset_sInter_iff]; simp lemma mem_closure_iff_forall_mem_flat (X : Set α) (hX : X ⊆ M.E := by aesop_mat) : e ∈ M.closure X ↔ ∀ F, M.Flat F → X ⊆ F → e ∈ F := by diff --git a/Mathlib/Data/PFunctor/Multivariate/M.lean b/Mathlib/Data/PFunctor/Multivariate/M.lean index cb6ed15fd11ac..3760a1050ad44 100644 --- a/Mathlib/Data/PFunctor/Multivariate/M.lean +++ b/Mathlib/Data/PFunctor/Multivariate/M.lean @@ -284,7 +284,7 @@ theorem M.bisim' {α : TypeVec n} (R : P.M α → P.M α → Prop) induction Hr · rw [← Quot.factor_mk_eq R (Relation.EqvGen R) this] rwa [appendFun_comp_id, ← MvFunctor.map_map, ← MvFunctor.map_map, h] - all_goals aesop + all_goals simp_all theorem M.dest_map {α β : TypeVec n} (g : α ⟹ β) (x : P.M α) : M.dest P (g <$$> x) = (appendFun g fun x => g <$$> x) <$$> M.dest P x := by diff --git a/Mathlib/Data/Part.lean b/Mathlib/Data/Part.lean index 842a941412357..98ebf824ed437 100644 --- a/Mathlib/Data/Part.lean +++ b/Mathlib/Data/Part.lean @@ -397,7 +397,7 @@ theorem assert_pos {p : Prop} {f : p → Part α} (h : p) : assert p f = f h := simp only [h', mk.injEq, h, exists_prop_of_true, true_and] apply Function.hfunext · simp only [h, h', exists_prop_of_true] - · aesop + · simp theorem assert_neg {p : Prop} {f : p → Part α} (h : ¬p) : assert p f = none := by dsimp [assert, none]; congr diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean index edc1b664e2ba8..45478ac73f241 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean @@ -82,7 +82,7 @@ noncomputable def coproductCofan : Cocone F where ι := { app := fun j => ⟨colimit.ι (C := SheafedSpace.{u+1, u, u} CommRingCatMax.{u, u}) (F ⋙ forgetToSheafedSpace) j, inferInstance⟩ - naturality := fun ⟨j⟩ ⟨j'⟩ ⟨⟨(f : j = j')⟩⟩ => by subst f; aesop } + naturality := fun ⟨j⟩ ⟨j'⟩ ⟨⟨(f : j = j')⟩⟩ => by subst f; simp } /-- The explicit coproduct cofan constructed in `coproduct_cofan` is indeed a colimit. -/ noncomputable def coproductCofanIsColimit : IsColimit (coproductCofan F) where diff --git a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean index e5a725fd32218..2ad3ed38cd7f4 100644 --- a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean +++ b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean @@ -113,7 +113,7 @@ lemma eq_finset_univ [Fintype G] [DecidableEq G] rw [card_four'] repeat rw [card_insert_of_not_mem] on_goal 4 => simpa using mul_not_mem_of_exponent_two (by simp) hx hy hxy - all_goals aesop + all_goals simp_all @[to_additive] lemma eq_mul_of_ne_all {x y z : G} (hx : x ≠ 1) diff --git a/Mathlib/LinearAlgebra/FreeProduct/Basic.lean b/Mathlib/LinearAlgebra/FreeProduct/Basic.lean index a6d1297916a19..15dd2e3d55456 100644 --- a/Mathlib/LinearAlgebra/FreeProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/FreeProduct/Basic.lean @@ -189,7 +189,7 @@ theorem mul_injections (a₁ a₂ : A i) : ι' R A (DirectSum.lof R I A i a₁) * ι' R A (DirectSum.lof R I A i a₂) = ι' R A (DirectSum.lof R I A i (a₁ * a₂)) := by convert RingQuot.mkAlgHom_rel R <| rel.prod - aesop + simp /--The `i`th canonical injection, from `A i` to the free product, as a linear map.-/ diff --git a/Mathlib/LinearAlgebra/InvariantBasisNumber.lean b/Mathlib/LinearAlgebra/InvariantBasisNumber.lean index fd7f09efc5db7..2a0d1fc752546 100644 --- a/Mathlib/LinearAlgebra/InvariantBasisNumber.lean +++ b/Mathlib/LinearAlgebra/InvariantBasisNumber.lean @@ -241,8 +241,8 @@ theorem nontrivial_of_invariantBasisNumber : Nontrivial R := by exact { toFun := 0 invFun := 0 - map_add' := by aesop - map_smul' := by aesop + map_add' := by simp + map_smul' := by simp left_inv := fun _ => by simp [eq_iff_true_of_subsingleton] right_inv := fun _ => by simp [eq_iff_true_of_subsingleton] } diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 5de11b6d8fc35..f4fa630b71b58 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -1081,8 +1081,8 @@ map from `Π i, M₁ᵢ ⟶ M₁ᵢ'` to `M ⟶ M₂` via `(fᵢ) ↦ v ↦ g(f MultilinearMap R M₁' M₂ →ₗ[R] MultilinearMap R (fun i ↦ M₁ i →ₗ[R] M₁' i) (MultilinearMap R M₁ M₂) where toFun g := (LinearMap.applyₗ g).compMultilinearMap compLinearMapMultilinear - map_add' := by aesop - map_smul' := by aesop + map_add' := by simp + map_smul' := by simp end diff --git a/Mathlib/LinearAlgebra/Semisimple.lean b/Mathlib/LinearAlgebra/Semisimple.lean index c7032628a6166..feb4bde80dc93 100644 --- a/Mathlib/LinearAlgebra/Semisimple.lean +++ b/Mathlib/LinearAlgebra/Semisimple.lean @@ -73,8 +73,7 @@ lemma isSemisimple_iff' : lemma isSemisimple_iff : f.IsSemisimple ↔ ∀ p ∈ invtSubmodule f, ∃ q ∈ invtSubmodule f, IsCompl p q := by - simp_rw [isSemisimple_iff'] - aesop + simp [isSemisimple_iff'] lemma isSemisimple_restrict_iff (p) (hp : p ∈ invtSubmodule f) : IsSemisimple (LinearMap.restrict f hp) ↔ diff --git a/Mathlib/LinearAlgebra/TensorProduct/Quotient.lean b/Mathlib/LinearAlgebra/TensorProduct/Quotient.lean index 21ad44e035faf..0c8559a3153a9 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Quotient.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Quotient.lean @@ -107,7 +107,7 @@ noncomputable def quotientTensorEquiv (m : Submodule R M) : erw [Submodule.map_id, Submodule.map_id] simp only [sup_eq_left] rw [map_range_eq_span_tmul, map_range_eq_span_tmul] - aesop) + simp) @[simp] lemma quotientTensorEquiv_apply_tmul_mk (m : Submodule R M) (x : M) (y : N) : @@ -137,7 +137,7 @@ noncomputable def tensorQuotientEquiv (n : Submodule R N) : erw [Submodule.map_id, Submodule.map_id] simp only [sup_eq_right] rw [map_range_eq_span_tmul, map_range_eq_span_tmul] - aesop) + simp) @[simp] lemma tensorQuotientEquiv_apply_mk_tmul (n : Submodule R N) (x : M) (y : N) : diff --git a/Mathlib/ModelTheory/ElementaryMaps.lean b/Mathlib/ModelTheory/ElementaryMaps.lean index 6b3b1c267638e..b044135204fda 100644 --- a/Mathlib/ModelTheory/ElementaryMaps.lean +++ b/Mathlib/ModelTheory/ElementaryMaps.lean @@ -138,14 +138,14 @@ theorem map_constants (φ : M ↪ₑ[L] N) (c : L.Constants) : φ c = c := def toEmbedding (f : M ↪ₑ[L] N) : M ↪[L] N where toFun := f inj' := f.injective - map_fun' {_} f x := by aesop - map_rel' {_} R x := by aesop + map_fun' {_} f x := by simp + map_rel' {_} R x := by simp /-- An elementary embedding is also a first-order homomorphism. -/ def toHom (f : M ↪ₑ[L] N) : M →[L] N where toFun := f - map_fun' {_} f x := by aesop - map_rel' {_} R x := by aesop + map_fun' {_} f x := by simp + map_rel' {_} R x := by simp @[simp] theorem toEmbedding_toHom (f : M ↪ₑ[L] N) : f.toEmbedding.toHom = f.toHom := @@ -281,7 +281,7 @@ namespace Equiv /-- A first-order equivalence is also an elementary embedding. -/ def toElementaryEmbedding (f : M ≃[L] N) : M ↪ₑ[L] N where toFun := f - map_formula' n φ x := by aesop + map_formula' n φ x := by simp @[simp] theorem toElementaryEmbedding_toEmbedding (f : M ≃[L] N) : diff --git a/Mathlib/ModelTheory/Substructures.lean b/Mathlib/ModelTheory/Substructures.lean index 8b6541e29a94e..6ff42480b1d52 100644 --- a/Mathlib/ModelTheory/Substructures.lean +++ b/Mathlib/ModelTheory/Substructures.lean @@ -890,8 +890,8 @@ noncomputable def substructureEquivMap (f : M ↪[L] N) (s : L.Substructure M) : (codRestrict (s.map f.toHom) (f.domRestrict s) (fun ⟨m, hm⟩ => ⟨m, hm, rfl⟩) ⟨m, hm⟩).2).2) right_inv := fun ⟨_, hn⟩ => Subtype.mk_eq_mk.2 (Classical.choose_spec hn).2 - map_fun' {n} f x := by aesop - map_rel' {n} R x := by aesop + map_fun' {n} f x := by simp + map_rel' {n} R x := by simp @[simp] theorem substructureEquivMap_apply (f : M ↪[L] N) (p : L.Substructure M) (x : p) : @@ -910,8 +910,8 @@ theorem subtype_substructureEquivMap (f : M ↪[L] N) (s : L.Substructure M) : left_inv m := f.injective (Classical.choose_spec (codRestrict f.toHom.range f f.toHom.mem_range_self m).2) right_inv := fun ⟨_, hn⟩ => Subtype.mk_eq_mk.2 (Classical.choose_spec hn) - map_fun' {n} f x := by aesop - map_rel' {n} R x := by aesop + map_fun' {n} f x := by simp + map_rel' {n} R x := by simp @[simp] theorem equivRange_apply (f : M ↪[L] N) (x : M) : (f.equivRange x : N) = f x := diff --git a/Mathlib/Order/BooleanSubalgebra.lean b/Mathlib/Order/BooleanSubalgebra.lean index 24751706b1099..4419c6cc58077 100644 --- a/Mathlib/Order/BooleanSubalgebra.lean +++ b/Mathlib/Order/BooleanSubalgebra.lean @@ -160,10 +160,10 @@ instance instTop : Top (BooleanSubalgebra α) where /-- The trivial boolean subalgebra of a lattice. -/ instance instBot : Bot (BooleanSubalgebra α) where bot.carrier := {⊥, ⊤} - bot.bot_mem' := by aesop - bot.compl_mem' := by aesop - bot.supClosed' _ := by aesop - bot.infClosed' _ := by aesop + bot.bot_mem' := by simp + bot.compl_mem' := by simp + bot.supClosed' _ := by simp + bot.infClosed' _ := by simp /-- The inf of two boolean subalgebras is their intersection. -/ instance instInf : Min (BooleanSubalgebra α) where diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 734ac6c8622b2..af74169d09b3b 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -426,7 +426,7 @@ theorem strong_law_aux1 {c : ℝ} (c_one : 1 < c) {ε : ℝ} (εpos : 0 < ε) : · simp only [mem_sigma, mem_range, filter_congr_decidable, mem_filter, and_imp, Sigma.forall] exact fun a b haN hb ↦ ⟨hb.trans_le <| u_mono <| Nat.le_pred_of_lt haN, haN, hb⟩ - all_goals aesop + all_goals simp _ ≤ ∑ j ∈ range (u (N - 1)), c ^ 5 * (c - 1)⁻¹ ^ 3 / ↑j ^ 2 * Var[Y j] := by apply sum_le_sum fun j hj => ?_ rcases @eq_zero_or_pos _ _ j with (rfl | hj) diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index ff50887857737..1c8844cd09b48 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -797,7 +797,7 @@ def shortComplexH1Iso : (inhomogeneousCochains A).sc' 0 1 2 ≅ shortComplexH1 A /-- The 1-cocycles of the complex of inhomogeneous cochains of `A` are isomorphic to `oneCocycles A`, which is a simpler type. -/ def isoOneCocycles : cocycles A 1 ≅ ModuleCat.of k (oneCocycles A) := - (inhomogeneousCochains A).cyclesIsoSc' _ _ _ (by aesop) (by aesop) ≪≫ + (inhomogeneousCochains A).cyclesIsoSc' _ _ _ (by simp) (by simp) ≪≫ cyclesMapIso (shortComplexH1Iso A) ≪≫ (shortComplexH1 A).moduleCatCyclesIso lemma isoOneCocycles_hom_comp_subtype : @@ -817,7 +817,7 @@ lemma toCocycles_comp_isoOneCocycles_hom : /-- The 1st group cohomology of `A`, defined as the 1st cohomology of the complex of inhomogeneous cochains, is isomorphic to `oneCocycles A ⧸ oneCoboundaries A`, which is a simpler type. -/ def isoH1 : groupCohomology A 1 ≅ ModuleCat.of k (H1 A) := - (inhomogeneousCochains A).homologyIsoSc' _ _ _ (by aesop) (by aesop) ≪≫ + (inhomogeneousCochains A).homologyIsoSc' _ _ _ (by simp) (by simp) ≪≫ homologyMapIso (shortComplexH1Iso A) ≪≫ (shortComplexH1 A).moduleCatHomologyIso lemma groupCohomologyπ_comp_isoH1_hom : @@ -846,7 +846,7 @@ def shortComplexH2Iso : /-- The 2-cocycles of the complex of inhomogeneous cochains of `A` are isomorphic to `twoCocycles A`, which is a simpler type. -/ def isoTwoCocycles : cocycles A 2 ≅ ModuleCat.of k (twoCocycles A) := - (inhomogeneousCochains A).cyclesIsoSc' _ _ _ (by aesop) (by aesop) ≪≫ + (inhomogeneousCochains A).cyclesIsoSc' _ _ _ (by simp) (by simp) ≪≫ cyclesMapIso (shortComplexH2Iso A) ≪≫ (shortComplexH2 A).moduleCatCyclesIso lemma isoTwoCocycles_hom_comp_subtype : @@ -866,7 +866,7 @@ lemma toCocycles_comp_isoTwoCocycles_hom : /-- The 2nd group cohomology of `A`, defined as the 2nd cohomology of the complex of inhomogeneous cochains, is isomorphic to `twoCocycles A ⧸ twoCoboundaries A`, which is a simpler type. -/ def isoH2 : groupCohomology A 2 ≅ ModuleCat.of k (H2 A) := - (inhomogeneousCochains A).homologyIsoSc' _ _ _ (by aesop) (by aesop) ≪≫ + (inhomogeneousCochains A).homologyIsoSc' _ _ _ (by simp) (by simp) ≪≫ homologyMapIso (shortComplexH2Iso A) ≪≫ (shortComplexH2 A).moduleCatHomologyIso lemma groupCohomologyπ_comp_isoH2_hom : diff --git a/Mathlib/RingTheory/AdicCompletion/Functoriality.lean b/Mathlib/RingTheory/AdicCompletion/Functoriality.lean index 4ac8d7b974ccc..80039d291af17 100644 --- a/Mathlib/RingTheory/AdicCompletion/Functoriality.lean +++ b/Mathlib/RingTheory/AdicCompletion/Functoriality.lean @@ -128,7 +128,7 @@ private theorem adicCompletionAux_val_apply (f : M →ₗ[R] N) {n : ℕ} (x : A def map (f : M →ₗ[R] N) : AdicCompletion I M →ₗ[AdicCompletion I R] AdicCompletion I N where toFun := adicCompletionAux I f - map_add' := by aesop + map_add' := by simp map_smul' r x := by ext n simp only [adicCompletionAux_val_apply, smul_eval, smul_eq_mul, RingHom.id_apply] diff --git a/Mathlib/Topology/Category/Profinite/AsLimit.lean b/Mathlib/Topology/Category/Profinite/AsLimit.lean index 72fb66b943303..a417cb6fdd284 100644 --- a/Mathlib/Topology/Category/Profinite/AsLimit.lean +++ b/Mathlib/Topology/Category/Profinite/AsLimit.lean @@ -45,7 +45,7 @@ def fintypeDiagram : DiscreteQuotient X ⥤ FintypeCat where map f := DiscreteQuotient.ofLE f.le -- Porting note: `map_comp` used to be proved by default by `aesop_cat`. -- once `aesop_cat` can prove this again, remove the entire `map_comp` here. - map_comp _ _ := by ext; aesop_cat + map_comp _ _ := by ext; simp /-- An abbreviation for `X.fintypeDiagram ⋙ FintypeCat.toProfinite`. -/ abbrev diagram : DiscreteQuotient X ⥤ Profinite := diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index 745a728ee1127..9147bdcc6b1f6 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -215,7 +215,7 @@ theorem union_mem_nhds_of_mem_nhdsWithin {b : α} {L : Set α} (hL : L ∈ nhdsWithin b I₁) {R : Set α} (hR : R ∈ nhdsWithin b I₂) : L ∪ R ∈ nhds b := by rw [← nhdsWithin_univ b, h, nhdsWithin_union] - exact ⟨mem_of_superset hL (by aesop), mem_of_superset hR (by aesop)⟩ + exact ⟨mem_of_superset hL (by simp), mem_of_superset hR (by simp)⟩ /-- Writing a punctured neighborhood filter as a sup of left and right filters. -/ diff --git a/Mathlib/Topology/Sheaves/SheafCondition/OpensLeCover.lean b/Mathlib/Topology/Sheaves/SheafCondition/OpensLeCover.lean index 318e4ad961054..d316325a6727d 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/OpensLeCover.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/OpensLeCover.lean @@ -154,7 +154,7 @@ def generateEquivalenceOpensLe (hY : Y = iSup U) : inverse := generateEquivalenceOpensLe_inverse' _ hY unitIso := eqToIso <| CategoryTheory.Functor.ext (by rintro ⟨⟨_, _⟩, _⟩; dsimp; congr) - (by intros; refine Over.OverMorphism.ext ?_; aesop_cat) + (by intros; refine Over.OverMorphism.ext ?_; simp) counitIso := eqToIso <| CategoryTheory.Functor.hext (by intro; refine FullSubcategory.ext ?_; rfl) (by intros; rfl) From 10a421547e97c1ef37eb44e7eccfd52664072929 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Mon, 6 Jan 2025 07:01:31 +0000 Subject: [PATCH 169/189] chore(NumberTheory/NumberField/AdeleRing): refactor adele rings (#20500) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The current `AdeleRing ℚ` explicitly involves `𝓞 ℚ` and is not much fun to work with. This PR enables `AdeleRing ℤ ℚ`. --- .../NumberTheory/NumberField/AdeleRing.lean | 43 +++++++++++-------- 1 file changed, 26 insertions(+), 17 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/AdeleRing.lean b/Mathlib/NumberTheory/NumberField/AdeleRing.lean index 1ee25ce15fd5f..927332bf8bdeb 100644 --- a/Mathlib/NumberTheory/NumberField/AdeleRing.lean +++ b/Mathlib/NumberTheory/NumberField/AdeleRing.lean @@ -41,8 +41,6 @@ open InfinitePlace AbsoluteValue.Completion InfinitePlace.Completion DedekindDom open scoped Classical -variable (K : Type*) [Field K] - /-! ## The infinite adele ring The infinite adele ring is the finite product of completions of a number field over its @@ -51,10 +49,12 @@ infinite places. See `NumberField.InfinitePlace` for the definition of an infini -/ /-- The infinite adele ring of a number field. -/ -def InfiniteAdeleRing := (v : InfinitePlace K) → v.Completion +def InfiniteAdeleRing (K : Type*) [Field K] := (v : InfinitePlace K) → v.Completion namespace InfiniteAdeleRing +variable (K : Type*) [Field K] + instance : CommRing (InfiniteAdeleRing K) := Pi.commRing instance : Inhabited (InfiniteAdeleRing K) := ⟨0⟩ @@ -115,38 +115,47 @@ theorem mixedEmbedding_eq_algebraMap_comp {x : K} : end InfiniteAdeleRing -variable [NumberField K] - /-! ## The adele ring -/ -/-- The adele ring of a number field. -/ -def AdeleRing := InfiniteAdeleRing K × FiniteAdeleRing (𝓞 K) K +/-- `AdeleRing (𝓞 K) K` is the adele ring of a number field `K`. + +More generally `AdeleRing R K` can be used if `K` is the field of fractions +of the Dedekind domain `R`. This enables use of rings like `AdeleRing ℤ ℚ`, which +in practice are easier to work with than `AdeleRing (𝓞 ℚ) ℚ`. + +Note that this definition does not give the correct answer in the function field case. +-/ +def AdeleRing (R K : Type*) [CommRing R] [IsDedekindDomain R] [Field K] + [Algebra R K] [IsFractionRing R K] := InfiniteAdeleRing K × FiniteAdeleRing R K namespace AdeleRing -instance : CommRing (AdeleRing K) := Prod.instCommRing +variable (R K : Type*) [CommRing R] [IsDedekindDomain R] [Field K] + [Algebra R K] [IsFractionRing R K] + +instance : CommRing (AdeleRing R K) := Prod.instCommRing -instance : Inhabited (AdeleRing K) := ⟨0⟩ +instance : Inhabited (AdeleRing R K) := ⟨0⟩ -instance : TopologicalSpace (AdeleRing K) := instTopologicalSpaceProd +instance : TopologicalSpace (AdeleRing R K) := instTopologicalSpaceProd -instance : TopologicalRing (AdeleRing K) := instTopologicalRingProd +instance : TopologicalRing (AdeleRing R K) := instTopologicalRingProd -instance : Algebra K (AdeleRing K) := Prod.algebra _ _ _ +instance : Algebra K (AdeleRing R K) := Prod.algebra _ _ _ @[simp] theorem algebraMap_fst_apply (x : K) (v : InfinitePlace K) : - (algebraMap K (AdeleRing K) x).1 v = x := rfl + (algebraMap K (AdeleRing R K) x).1 v = x := rfl @[simp] -theorem algebraMap_snd_apply (x : K) (v : HeightOneSpectrum (𝓞 K)) : - (algebraMap K (AdeleRing K) x).2 v = x := rfl +theorem algebraMap_snd_apply (x : K) (v : HeightOneSpectrum R) : + (algebraMap K (AdeleRing R K) x).2 v = x := rfl -theorem algebraMap_injective : Function.Injective (algebraMap K (AdeleRing K)) := +theorem algebraMap_injective [NumberField K] : Function.Injective (algebraMap K (AdeleRing R K)) := fun _ _ hxy => (algebraMap K _).injective (Prod.ext_iff.1 hxy).1 /-- The subgroup of principal adeles `(x)ᵥ` where `x ∈ K`. -/ -abbrev principalSubgroup : AddSubgroup (AdeleRing K) := (algebraMap K _).range.toAddSubgroup +abbrev principalSubgroup : AddSubgroup (AdeleRing R K) := (algebraMap K _).range.toAddSubgroup end AdeleRing From 9afdff78cbdf653902190167fd1429e4adc4ac89 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 6 Jan 2025 07:44:01 +0000 Subject: [PATCH 170/189] chore(LinearIndependent): generalize to semirings (#20480) After the definition is changed in #18426, many results can now be generalized to semirings. Co-authored-by: Eric Wieser --- Mathlib/Data/Finsupp/Basic.lean | 26 +- Mathlib/Data/Set/Function.lean | 4 + Mathlib/LinearAlgebra/Dimension/Basic.lean | 2 +- .../Finsupp/LinearCombination.lean | 5 + Mathlib/LinearAlgebra/Finsupp/Supported.lean | 6 +- Mathlib/LinearAlgebra/LinearIndependent.lean | 1274 +++++++++-------- 6 files changed, 735 insertions(+), 582 deletions(-) diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index 8f9546c197a70..675105181e1fc 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -905,17 +905,24 @@ theorem subtypeDomain_apply {a : Subtype p} {v : α →₀ M} : (subtypeDomain p theorem subtypeDomain_zero : subtypeDomain p (0 : α →₀ M) = 0 := rfl -theorem subtypeDomain_eq_zero_iff' {f : α →₀ M} : f.subtypeDomain p = 0 ↔ ∀ x, p x → f x = 0 := by - classical simp_rw [← support_eq_empty, support_subtypeDomain, subtype_eq_empty, - not_mem_support_iff] +theorem subtypeDomain_eq_iff_forall {f g : α →₀ M} : + f.subtypeDomain p = g.subtypeDomain p ↔ ∀ x, p x → f x = g x := by + simp_rw [DFunLike.ext_iff, subtypeDomain_apply, Subtype.forall] + +theorem subtypeDomain_eq_iff {f g : α →₀ M} + (hf : ∀ x ∈ f.support, p x) (hg : ∀ x ∈ g.support, p x) : + f.subtypeDomain p = g.subtypeDomain p ↔ f = g := + subtypeDomain_eq_iff_forall.trans + ⟨fun H ↦ Finsupp.ext fun _a ↦ (em _).elim (H _ <| hf _ ·) fun haf ↦ (em _).elim (H _ <| hg _ ·) + fun hag ↦ (not_mem_support_iff.mp haf).trans (not_mem_support_iff.mp hag).symm, + fun H _ _ ↦ congr($H _)⟩ + +theorem subtypeDomain_eq_zero_iff' {f : α →₀ M} : f.subtypeDomain p = 0 ↔ ∀ x, p x → f x = 0 := + subtypeDomain_eq_iff_forall (g := 0) theorem subtypeDomain_eq_zero_iff {f : α →₀ M} (hf : ∀ x ∈ f.support, p x) : f.subtypeDomain p = 0 ↔ f = 0 := - subtypeDomain_eq_zero_iff'.trans - ⟨fun H => - ext fun x => by - classical exact if hx : p x then H x hx else not_mem_support_iff.1 <| mt (hf x) hx, - fun H x _ => by simp [H]⟩ + subtypeDomain_eq_iff (g := 0) hf (by simp) @[to_additive] theorem prod_subtypeDomain_index [CommMonoid N] {v : α →₀ M} {h : α → M → N} @@ -1504,7 +1511,8 @@ end /-- Given an `AddCommMonoid M` and `s : Set α`, `restrictSupportEquiv s M` is the `Equiv` between the subtype of finitely supported functions with support contained in `s` and the type of finitely supported functions from `s`. -/ -def restrictSupportEquiv (s : Set α) (M : Type*) [AddCommMonoid M] : +-- TODO: add [DecidablePred (· ∈ s)] as an assumption +@[simps] def restrictSupportEquiv (s : Set α) (M : Type*) [AddCommMonoid M] : { f : α →₀ M // ↑f.support ⊆ s } ≃ (s →₀ M) where toFun f := subtypeDomain (· ∈ s) f.1 invFun f := letI := Classical.decPred (· ∈ s); ⟨f.extendDomain, support_extendDomain_subset _⟩ diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index f2c6f67261d56..c29e123ed9246 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -529,6 +529,10 @@ theorem _root_.Function.Injective.injOn_range (h : Injective (g ∘ f)) : InjOn rintro _ ⟨x, rfl⟩ _ ⟨y, rfl⟩ H exact congr_arg f (h H) +theorem _root_.Set.InjOn.injective_iff (s : Set β) (h : InjOn g s) (hs : range f ⊆ s) : + Injective (g ∘ f) ↔ Injective f := + ⟨(·.of_comp), fun h _ ↦ by aesop⟩ + theorem injOn_iff_injective : InjOn f s ↔ Injective (s.restrict f) := ⟨fun H a b h => Subtype.eq <| H a.2 b.2 h, fun H a as b bs h => congr_arg Subtype.val <| @H ⟨a, as⟩ ⟨b, bs⟩ h⟩ diff --git a/Mathlib/LinearAlgebra/Dimension/Basic.lean b/Mathlib/LinearAlgebra/Dimension/Basic.lean index 5e63e9aa3eb1a..124960fd5f2ea 100644 --- a/Mathlib/LinearAlgebra/Dimension/Basic.lean +++ b/Mathlib/LinearAlgebra/Dimension/Basic.lean @@ -68,7 +68,7 @@ end namespace LinearIndependent -variable [Ring R] [AddCommGroup M] [Module R M] +variable [Semiring R] [AddCommMonoid M] [Module R M] variable [Nontrivial R] diff --git a/Mathlib/LinearAlgebra/Finsupp/LinearCombination.lean b/Mathlib/LinearAlgebra/Finsupp/LinearCombination.lean index 24ebf16e2af46..c1b6ee8d30eb5 100644 --- a/Mathlib/LinearAlgebra/Finsupp/LinearCombination.lean +++ b/Mathlib/LinearAlgebra/Finsupp/LinearCombination.lean @@ -283,6 +283,11 @@ theorem linearCombinationOn_range (s : Set α) : @[deprecated (since := "2024-08-29")] alias totalOn_range := linearCombinationOn_range +theorem linearCombination_restrict (s : Set α) : + linearCombination R (s.restrict v) = Submodule.subtype _ ∘ₗ + linearCombinationOn α M R v s ∘ₗ (supportedEquivFinsupp s).symm.toLinearMap := by + classical ext; simp [linearCombinationOn] + theorem linearCombination_comp (f : α' → α) : linearCombination R (v ∘ f) = (linearCombination R v).comp (lmapDomain R R f) := by ext diff --git a/Mathlib/LinearAlgebra/Finsupp/Supported.lean b/Mathlib/LinearAlgebra/Finsupp/Supported.lean index db7332ff5630f..d779324d7594f 100644 --- a/Mathlib/LinearAlgebra/Finsupp/Supported.lean +++ b/Mathlib/LinearAlgebra/Finsupp/Supported.lean @@ -168,7 +168,7 @@ theorem disjoint_supported_supported_iff [Nontrivial M] {s t : Set α} : /-- Interpret `Finsupp.restrictSupportEquiv` as a linear equivalence between `supported M R s` and `s →₀ M`. -/ -def supportedEquivFinsupp (s : Set α) : supported M R s ≃ₗ[R] s →₀ M := by +@[simps!] def supportedEquivFinsupp (s : Set α) : supported M R s ≃ₗ[R] s →₀ M := by let F : supported M R s ≃ (s →₀ M) := restrictSupportEquiv s M refine F.toLinearEquiv ?_ have : @@ -178,6 +178,10 @@ def supportedEquivFinsupp (s : Set α) : supported M R s ≃ₗ[R] s →₀ M := rw [this] exact LinearMap.isLinear _ +@[simp] theorem supportedEquivFinsupp_symm_apply_coe (s : Set α) [DecidablePred (· ∈ s)] + (f : s →₀ M) : (supportedEquivFinsupp (R := R) s).symm f = f.extendDomain := by + convert restrictSupportEquiv_symm_apply_coe .. + section LMapDomain variable {α' : Type*} {α'' : Type*} (M R) diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index 973a47eec399b..561a2be1ce8fa 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -4,8 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Alexander Bentkamp, Anne Baanen -/ import Mathlib.Algebra.BigOperators.Fin +import Mathlib.Data.Fin.Tuple.Reflection import Mathlib.Data.Set.Subsingleton import Mathlib.Lean.Expr.ExtraRecognizers +import Mathlib.LinearAlgebra.Finsupp.SumProd import Mathlib.LinearAlgebra.Prod import Mathlib.Tactic.FinCases import Mathlib.Tactic.LinearCombination @@ -101,7 +103,7 @@ variable [Module R M] [Module R M'] variable (R) (v) /-- `LinearIndependent R v` states the family of vectors `v` is linearly independent over `R`. -/ def LinearIndependent : Prop := - Function.Injective (Finsupp.linearCombination R v) + Injective (Finsupp.linearCombination R v) open Lean PrettyPrinter.Delaborator SubExpr in /-- Delaborator for `LinearIndependent` that suggests pretty printing with type hints @@ -126,46 +128,678 @@ def delabLinearIndependent : Delab := variable {R v} theorem linearIndependent_iff_injective_linearCombination : - LinearIndependent R v ↔ Function.Injective (Finsupp.linearCombination R v) := Iff.rfl + LinearIndependent R v ↔ Injective (Finsupp.linearCombination R v) := Iff.rfl alias ⟨LinearIndependent.injective_linearCombination, _⟩ := linearIndependent_iff_injective_linearCombination +theorem Fintype.linearIndependent_iff_injective [Fintype ι] : + LinearIndependent R v ↔ Injective (Fintype.linearCombination R ℕ v) := by + simp [← Finsupp.linearCombination_eq_fintype_linearCombination, LinearIndependent] + +theorem LinearIndependent.injective [Nontrivial R] (hv : LinearIndependent R v) : Injective v := by + simpa [comp_def] + using Injective.comp hv (Finsupp.single_left_injective one_ne_zero) + theorem LinearIndependent.ne_zero [Nontrivial R] (i : ι) (hv : LinearIndependent R v) : v i ≠ 0 := by intro h - have := @hv (Finsupp.single i 1 : ι →₀ R) 0 ?_ - · simp at this + have := @hv (Finsupp.single i 1 : ι →₀ R) 0 (by simpa using h) + simp at this + +theorem linearIndependent_empty_type [IsEmpty ι] : LinearIndependent R v := + injective_of_subsingleton _ + +variable (R M) in +theorem linearIndependent_empty : LinearIndependent R (fun x => x : (∅ : Set M) → M) := + linearIndependent_empty_type + +/-- A subfamily of a linearly independent family (i.e., a composition with an injective map) is a +linearly independent family. -/ +theorem LinearIndependent.comp (h : LinearIndependent R v) (f : ι' → ι) (hf : Injective f) : + LinearIndependent R (v ∘ f) := by + simpa [comp_def] using Injective.comp h (Finsupp.mapDomain_injective hf) + +/-- A set of linearly independent vectors in a module `M` over a semiring `K` is also linearly +independent over a subring `R` of `K`. +The implementation uses minimal assumptions about the relationship between `R`, `K` and `M`. +The version where `K` is an `R`-algebra is `LinearIndependent.restrict_scalars_algebras`. -/ +theorem LinearIndependent.restrict_scalars [Semiring K] [SMulWithZero R K] [Module K M] + [IsScalarTower R K M] (hinj : Injective fun r : R ↦ r • (1 : K)) + (li : LinearIndependent K v) : LinearIndependent R v := by + intro x y hxy + let f := fun r : R => r • (1 : K) + have := @li (x.mapRange f (by simp [f])) (y.mapRange f (by simp [f])) ?_ + · ext i + exact hinj congr($this i) + simpa [Finsupp.linearCombination, f, Finsupp.sum_mapRange_index] + +@[deprecated (since := "2024-08-29")] alias linearIndependent_iff_injective_total := + linearIndependent_iff_injective_linearCombination + +@[deprecated (since := "2024-08-29")] alias LinearIndependent.injective_total := + LinearIndependent.injective_linearCombination + +-- This version makes `l₁` and `l₂` explicit. +theorem linearIndependent_iffₛ : + LinearIndependent R v ↔ + ∀ l₁ l₂, Finsupp.linearCombination R v l₁ = Finsupp.linearCombination R v l₂ → l₁ = l₂ := + Iff.rfl + +open Finset in +theorem linearIndependent_iff'ₛ : + LinearIndependent R v ↔ + ∀ s : Finset ι, ∀ f g : ι → R, ∑ i ∈ s, f i • v i = ∑ i ∈ s, g i • v i → ∀ i ∈ s, f i = g i := + linearIndependent_iffₛ.trans + ⟨fun hv s f g eq i his ↦ by + have h := + hv (∑ i ∈ s, Finsupp.single i (f i)) (∑ i ∈ s, Finsupp.single i (g i)) <| by + simpa only [map_sum, Finsupp.linearCombination_single] using eq + have (f : ι → R) : f i = (∑ j ∈ s, Finsupp.single j (f j)) i := + calc + f i = (Finsupp.lapply i : (ι →₀ R) →ₗ[R] R) (Finsupp.single i (f i)) := by + { rw [Finsupp.lapply_apply, Finsupp.single_eq_same] } + _ = ∑ j ∈ s, (Finsupp.lapply i : (ι →₀ R) →ₗ[R] R) (Finsupp.single j (f j)) := + Eq.symm <| + Finset.sum_eq_single i + (fun j _hjs hji => by rw [Finsupp.lapply_apply, Finsupp.single_eq_of_ne hji]) + fun hnis => hnis.elim his + _ = (∑ j ∈ s, Finsupp.single j (f j)) i := (map_sum ..).symm + rw [this f, this g, h], + fun hv f g hl ↦ + Finsupp.ext fun _ ↦ by + classical + refine _root_.by_contradiction fun hni ↦ hni <| hv (f.support ∪ g.support) f g ?_ _ ?_ + · rwa [← sum_subset subset_union_left, ← sum_subset subset_union_right] <;> + rintro i - hi <;> rw [Finsupp.not_mem_support_iff.mp hi, zero_smul] + · contrapose! hni + simp_rw [not_mem_union, Finsupp.not_mem_support_iff] at hni + rw [hni.1, hni.2]⟩ + +theorem linearIndependent_iff''ₛ : + LinearIndependent R v ↔ + ∀ (s : Finset ι) (f g : ι → R), (∀ i ∉ s, f i = g i) → + ∑ i ∈ s, f i • v i = ∑ i ∈ s, g i • v i → ∀ i, f i = g i := by + classical + exact linearIndependent_iff'ₛ.trans + ⟨fun H s f g eq hv i ↦ if his : i ∈ s then H s f g hv i his else eq i his, + fun H s f g eq i hi ↦ by + convert + H s (fun j ↦ if j ∈ s then f j else 0) (fun j ↦ if j ∈ s then g j else 0) + (fun j hj ↦ (if_neg hj).trans (if_neg hj).symm) + (by simp_rw [ite_smul, zero_smul, Finset.sum_extend_by_zero, eq]) i <;> + exact (if_pos hi).symm⟩ + +theorem not_linearIndependent_iffₛ : + ¬LinearIndependent R v ↔ ∃ s : Finset ι, + ∃ f g : ι → R, ∑ i ∈ s, f i • v i = ∑ i ∈ s, g i • v i ∧ ∃ i ∈ s, f i ≠ g i := by + rw [linearIndependent_iff'ₛ] + simp only [exists_prop, not_forall] + +theorem Fintype.linearIndependent_iffₛ [Fintype ι] : + LinearIndependent R v ↔ ∀ f g : ι → R, ∑ i, f i • v i = ∑ i, g i • v i → ∀ i, f i = g i := by + simp_rw [Fintype.linearIndependent_iff_injective, + Injective, Fintype.linearCombination_apply, funext_iff] + +/-- A finite family of vectors `v i` is linear independent iff the linear map that sends +`c : ι → R` to `∑ i, c i • v i` is injective. -/ +theorem Fintype.linearIndependent_iff'ₛ [Fintype ι] [DecidableEq ι] : + LinearIndependent R v ↔ + Injective (LinearMap.lsum R (fun _ ↦ R) ℕ fun i ↦ LinearMap.id.smulRight (v i)) := by + simp [Fintype.linearIndependent_iffₛ, Injective, funext_iff] + +theorem Fintype.not_linearIndependent_iffₛ [Fintype ι] : + ¬LinearIndependent R v ↔ ∃ f g : ι → R, ∑ i, f i • v i = ∑ i, g i • v i ∧ ∃ i, f i ≠ g i := by + simpa using not_iff_not.2 Fintype.linearIndependent_iffₛ + +lemma LinearIndependent.pair_iffₛ {x y : M} : + LinearIndependent R ![x, y] ↔ + ∀ (s t s' t' : R), s • x + t • y = s' • x + t' • y → s = s' ∧ t = t' := by + simp [Fintype.linearIndependent_iffₛ, Fin.forall_fin_two, ← FinVec.forall_iff]; rfl + +lemma LinearIndependent.eq_of_pair {x y : M} (h : LinearIndependent R ![x, y]) + {s t s' t' : R} (h' : s • x + t • y = s' • x + t' • y) : s = s' ∧ t = t' := + pair_iffₛ.mp h _ _ _ _ h' + +lemma LinearIndependent.eq_zero_of_pair' {x y : M} (h : LinearIndependent R ![x, y]) + {s t : R} (h' : s • x = t • y) : s = 0 ∧ t = 0 := by + suffices H : s = 0 ∧ 0 = t from ⟨H.1, H.2.symm⟩ + exact h.eq_of_pair (by simpa using h') + +lemma LinearIndependent.eq_zero_of_pair {x y : M} (h : LinearIndependent R ![x, y]) + {s t : R} (h' : s • x + t • y = 0) : s = 0 ∧ t = 0 := by + replace h := @h (.single 0 s + .single 1 t) 0 ?_ + · exact ⟨by simpa using congr($h 0), by simpa using congr($h 1)⟩ + simpa + +/-- A family is linearly independent if and only if all of its finite subfamily is +linearly independent. -/ +theorem linearIndependent_iff_finset_linearIndependent : + LinearIndependent R v ↔ ∀ (s : Finset ι), LinearIndependent R (v ∘ (Subtype.val : s → ι)) := + ⟨fun H _ ↦ H.comp _ Subtype.val_injective, fun H ↦ linearIndependent_iff'ₛ.2 fun s f g eq i hi ↦ + Fintype.linearIndependent_iffₛ.1 (H s) (f ∘ Subtype.val) (g ∘ Subtype.val) + (by simpa only [← s.sum_coe_sort] using eq) ⟨i, hi⟩⟩ + +theorem LinearIndependent.coe_range (i : LinearIndependent R v) : + LinearIndependent R ((↑) : range v → M) := by simpa using i.comp _ (rangeSplitting_injective v) + +/-- If `v` is an injective family of vectors such that `f ∘ v` is linearly independent, then `v` + spans a submodule disjoint from the kernel of `f` -/ +theorem Submodule.range_ker_disjoint {f : M →ₗ[R] M'} + (hv : LinearIndependent R (f ∘ v)) : + Disjoint (span R (range v)) (LinearMap.ker f) := by + rw [LinearIndependent, Finsupp.linearCombination_linear_comp] at hv + rw [disjoint_iff_inf_le, ← Set.image_univ, Finsupp.span_image_eq_map_linearCombination, + map_inf_eq_map_inf_comap, (LinearMap.ker_comp _ _).symm.trans + (LinearMap.ker_eq_bot_of_injective hv), inf_bot_eq, map_bot] + +/-- If `M / R` and `M' / R'` are modules, `i : R' → R` is a map, `j : M →+ M'` is a monoid map, +such that they are both injective, and compatible with the scalar +multiplications on `M` and `M'`, then `j` sends linearly independent families of vectors to +linearly independent families of vectors. As a special case, taking `R = R'` +it is `LinearIndependent.map'`. -/ +theorem LinearIndependent.map_of_injective_injectiveₛ {R' M' : Type*} + [Semiring R'] [AddCommMonoid M'] [Module R' M'] (hv : LinearIndependent R v) + (i : R' → R) (j : M →+ M') (hi : Function.Injective i) (hj : Function.Injective j) + (hc : ∀ (r : R') (m : M), j (i r • m) = r • j m) : LinearIndependent R' (j ∘ v) := by + rw [linearIndependent_iff'ₛ] at hv ⊢ + intro S r₁ r₂ H s hs + simp_rw [comp_apply, ← hc, ← map_sum] at H + exact hi <| hv _ _ _ (hj H) s hs + +/-- If the image of a family of vectors under a linear map is linearly independent, then so is +the original family. -/ +theorem LinearIndependent.of_comp (f : M →ₗ[R] M') (hfv : LinearIndependent R (f ∘ v)) : + LinearIndependent R v := by + rw [LinearIndependent, Finsupp.linearCombination_linear_comp, LinearMap.coe_comp] at hfv + exact hfv.of_comp + +/-- If `f` is a linear map injective on the span of the range of `v`, then the family `f ∘ v` +is linearly independent if and only if the family `v` is linearly independent. +See `LinearMap.linearIndependent_iff_of_disjoint` for the version with `Set.InjOn` replaced +by `Disjoint` when working over a ring. -/ +protected theorem LinearMap.linearIndependent_iff_of_injOn (f : M →ₗ[R] M') + (hf_inj : Set.InjOn f (span R (Set.range v))) : + LinearIndependent R (f ∘ v) ↔ LinearIndependent R v := by + simp_rw [LinearIndependent, Finsupp.linearCombination_linear_comp, coe_comp] + rw [hf_inj.injective_iff] + rw [← Finsupp.range_linearCombination, LinearMap.range_coe] + +/-- If a linear map is injective on the span of a family of linearly independent vectors, then +the family stays linearly independent after composing with the linear map. +See `LinearIndependent.map` for the version with `Set.InjOn` replaced by `Disjoint` +when working over a ring. -/ +theorem LinearIndependent.map_injOn (hv : LinearIndependent R v) (f : M →ₗ[R] M') + (hf_inj : Set.InjOn f (span R (Set.range v))) : LinearIndependent R (f ∘ v) := + (f.linearIndependent_iff_of_injOn hf_inj).mpr hv + +@[nontriviality] +theorem linearIndependent_of_subsingleton [Subsingleton R] : LinearIndependent R v := + linearIndependent_iffₛ.2 fun _l _l' _hl => Subsingleton.elim _ _ + +theorem linearIndependent_equiv (e : ι ≃ ι') {f : ι' → M} : + LinearIndependent R (f ∘ e) ↔ LinearIndependent R f := + ⟨fun h ↦ comp_id f ▸ e.self_comp_symm ▸ h.comp _ e.symm.injective, + fun h ↦ h.comp _ e.injective⟩ + +theorem linearIndependent_equiv' (e : ι ≃ ι') {f : ι' → M} {g : ι → M} (h : f ∘ e = g) : + LinearIndependent R g ↔ LinearIndependent R f := + h ▸ linearIndependent_equiv e + +theorem linearIndependent_subtype_range {ι} {f : ι → M} (hf : Injective f) : + LinearIndependent R ((↑) : range f → M) ↔ LinearIndependent R f := + Iff.symm <| linearIndependent_equiv' (Equiv.ofInjective f hf) rfl + +alias ⟨LinearIndependent.of_subtype_range, _⟩ := linearIndependent_subtype_range + +theorem LinearIndependent.to_subtype_range {ι} {f : ι → M} (hf : LinearIndependent R f) : + LinearIndependent R ((↑) : range f → M) := by + nontriviality R + exact (linearIndependent_subtype_range hf.injective).2 hf + +theorem LinearIndependent.to_subtype_range' {ι} {f : ι → M} (hf : LinearIndependent R f) {t} + (ht : range f = t) : LinearIndependent R ((↑) : t → M) := + ht ▸ hf.to_subtype_range + +theorem LinearIndependent.image_of_comp {ι ι'} (s : Set ι) (f : ι → ι') (g : ι' → M) + (hs : LinearIndependent R fun x : s => g (f x)) : + LinearIndependent R fun x : f '' s => g x := by + nontriviality R + have : InjOn f s := injOn_iff_injective.2 hs.injective.of_comp + exact (linearIndependent_equiv' (Equiv.Set.imageOfInjOn f s this) rfl).1 hs + +theorem LinearIndependent.image {ι} {s : Set ι} {f : ι → M} + (hs : LinearIndependent R fun x : s => f x) : + LinearIndependent R fun x : f '' s => (x : M) := by + convert LinearIndependent.image_of_comp s f id hs + +theorem LinearIndependent.group_smul {G : Type*} [hG : Group G] [DistribMulAction G R] + [DistribMulAction G M] [IsScalarTower G R M] [SMulCommClass G R M] {v : ι → M} + (hv : LinearIndependent R v) (w : ι → G) : LinearIndependent R (w • v) := by + rw [linearIndependent_iff''ₛ] at hv ⊢ + intro s g₁ g₂ hgs hsum i + refine (Group.isUnit (w i)).smul_left_cancel.mp ?_ + refine hv s (fun i ↦ w i • g₁ i) (fun i ↦ w i • g₂ i) (fun i hi ↦ ?_) ?_ i + · simp_rw [hgs i hi] + · simpa only [smul_assoc, smul_comm] using hsum + +-- This lemma cannot be proved with `LinearIndependent.group_smul` since the action of +-- `Rˣ` on `R` is not commutative. +theorem LinearIndependent.units_smul {v : ι → M} (hv : LinearIndependent R v) (w : ι → Rˣ) : + LinearIndependent R (w • v) := by + rw [linearIndependent_iff''ₛ] at hv ⊢ + intro s g₁ g₂ hgs hsum i + rw [← (w i).mul_left_inj] + refine hv s (fun i ↦ g₁ i • w i) (fun i ↦ g₂ i • w i) (fun i hi ↦ ?_) ?_ i + · simp_rw [hgs i hi] + · simpa only [smul_eq_mul, mul_smul, Pi.smul_apply'] using hsum + +theorem linearIndependent_image {ι} {s : Set ι} {f : ι → M} (hf : Set.InjOn f s) : + (LinearIndependent R fun x : s ↦ f x) ↔ LinearIndependent R fun x : f '' s => (x : M) := + linearIndependent_equiv' (Equiv.Set.imageOfInjOn _ _ hf) rfl + +theorem linearIndependent_span (hs : LinearIndependent R v) : + LinearIndependent R (M := span R (range v)) + (fun i : ι ↦ ⟨v i, subset_span (mem_range_self i)⟩) := + LinearIndependent.of_comp (span R (range v)).subtype hs + +/-- Every finite subset of a linearly independent set is linearly independent. -/ +theorem linearIndependent_finset_map_embedding_subtype (s : Set M) + (li : LinearIndependent R ((↑) : s → M)) (t : Finset s) : + LinearIndependent R ((↑) : Finset.map (Embedding.subtype s) t → M) := + li.comp (fun _ ↦ ⟨_, _⟩) <| by intro; aesop + +section Subtype + +/-! The following lemmas use the subtype defined by a set in `M` as the index set `ι`. -/ + +theorem linearIndependent_comp_subtypeₛ {s : Set ι} : + LinearIndependent R (v ∘ (↑) : s → M) ↔ + ∀ f ∈ Finsupp.supported R R s, ∀ g ∈ Finsupp.supported R R s, + Finsupp.linearCombination R v f = Finsupp.linearCombination R v g → f = g := by + simp only [linearIndependent_iffₛ, (· ∘ ·), Finsupp.mem_supported, + Finsupp.linearCombination_apply, Set.subset_def, Finset.mem_coe] + refine ⟨fun h l₁ h₁ l₂ h₂ eq ↦ (Finsupp.subtypeDomain_eq_iff h₁ h₂).1 <| h _ _ <| + (Finsupp.sum_subtypeDomain_index h₁).trans eq ▸ (Finsupp.sum_subtypeDomain_index h₂).symm, + fun h l₁ l₂ eq ↦ ?_⟩ + refine Finsupp.embDomain_injective (Embedding.subtype s) <| h _ ?_ _ ?_ ?_ + iterate 2 simpa using fun _ h _ ↦ h + simp_rw [Finsupp.embDomain_eq_mapDomain] + rwa [Finsupp.sum_mapDomain_index, Finsupp.sum_mapDomain_index] <;> + intros <;> simp only [zero_smul, add_smul] + +theorem linearDependent_comp_subtype'ₛ {s : Set ι} : + ¬LinearIndependent R (v ∘ (↑) : s → M) ↔ + ∃ f g : ι →₀ R, f ∈ Finsupp.supported R R s ∧ g ∈ Finsupp.supported R R s ∧ + Finsupp.linearCombination R v f = Finsupp.linearCombination R v g ∧ f ≠ g := by + simp [linearIndependent_comp_subtypeₛ, and_left_comm] + +/-- A version of `linearDependent_comp_subtype'ₛ` with `Finsupp.linearCombination` unfolded. -/ +theorem linearDependent_comp_subtypeₛ {s : Set ι} : + ¬LinearIndependent R (v ∘ (↑) : s → M) ↔ + ∃ f g : ι →₀ R, f ∈ Finsupp.supported R R s ∧ g ∈ Finsupp.supported R R s ∧ + ∑ i ∈ f.support, f i • v i = ∑ i ∈ g.support, g i • v i ∧ f ≠ g := + linearDependent_comp_subtype'ₛ + +theorem linearIndependent_subtypeₛ {s : Set M} : + LinearIndependent R (fun x ↦ x : s → M) ↔ + ∀ f ∈ Finsupp.supported R R s, ∀ g ∈ Finsupp.supported R R s, + Finsupp.linearCombination R id f = Finsupp.linearCombination R id g → f = g := + linearIndependent_comp_subtypeₛ (v := id) + +theorem linearIndependent_restrict_iff {s : Set ι} : + LinearIndependent R (s.restrict v) ↔ + Injective (Finsupp.linearCombinationOn ι M R v s) := by + simp [LinearIndependent, Finsupp.linearCombination_restrict] + +theorem linearIndependent_iff_linearCombinationOnₛ {s : Set M} : + LinearIndependent R (fun x ↦ x : s → M) ↔ + Injective (Finsupp.linearCombinationOn M M R id s) := + linearIndependent_restrict_iff (v := id) + +theorem LinearIndependent.restrict_of_comp_subtype {s : Set ι} + (hs : LinearIndependent R (v ∘ (↑) : s → M)) : LinearIndependent R (s.restrict v) := + hs + +theorem LinearIndependent.mono {t s : Set M} (h : t ⊆ s) + (hs : LinearIndependent R (fun x ↦ x : s → M)) : LinearIndependent R (fun x ↦ x : t → M) := + hs.comp _ (Set.inclusion_injective h) + +theorem linearIndependent_of_finite (s : Set M) + (H : ∀ t ⊆ s, Set.Finite t → LinearIndependent R (fun x ↦ x : t → M)) : + LinearIndependent R (fun x ↦ x : s → M) := + linearIndependent_subtypeₛ.2 fun f hf g hg eq ↦ + linearIndependent_subtypeₛ.1 (H _ (union_subset hf hg) <| (Finset.finite_toSet _).union <| + Finset.finite_toSet _) f Set.subset_union_left g Set.subset_union_right eq + +theorem linearIndependent_iUnion_of_directed {η : Type*} {s : η → Set M} (hs : Directed (· ⊆ ·) s) + (h : ∀ i, LinearIndependent R (fun x ↦ x : s i → M)) : + LinearIndependent R (fun x ↦ x : (⋃ i, s i) → M) := by + by_cases hη : Nonempty η + · refine linearIndependent_of_finite (⋃ i, s i) fun t ht ft => ?_ + rcases finite_subset_iUnion ft ht with ⟨I, fi, hI⟩ + rcases hs.finset_le fi.toFinset with ⟨i, hi⟩ + exact (h i).mono (Subset.trans hI <| iUnion₂_subset fun j hj => hi j (fi.mem_toFinset.2 hj)) + · refine (linearIndependent_empty R M).mono (t := iUnion (s ·)) ?_ + rintro _ ⟨_, ⟨i, _⟩, _⟩ + exact hη ⟨i⟩ + +theorem linearIndependent_sUnion_of_directed {s : Set (Set M)} (hs : DirectedOn (· ⊆ ·) s) + (h : ∀ a ∈ s, LinearIndependent R ((↑) : ((a : Set M) : Type _) → M)) : + LinearIndependent R (fun x => x : ⋃₀ s → M) := by + rw [sUnion_eq_iUnion] + exact linearIndependent_iUnion_of_directed hs.directed_val (by simpa using h) + +theorem linearIndependent_biUnion_of_directed {η} {s : Set η} {t : η → Set M} + (hs : DirectedOn (t ⁻¹'o (· ⊆ ·)) s) (h : ∀ a ∈ s, LinearIndependent R (fun x ↦ x : t a → M)) : + LinearIndependent R (fun x ↦ x : (⋃ a ∈ s, t a) → M) := by + rw [biUnion_eq_iUnion] + exact + linearIndependent_iUnion_of_directed (directed_comp.2 <| hs.directed_val) (by simpa using h) + +end Subtype + +/-- Linear independent families are injective, even if you multiply either side. -/ +theorem LinearIndependent.eq_of_smul_apply_eq_smul_apply {M : Type*} [AddCommMonoid M] [Module R M] + {v : ι → M} (li : LinearIndependent R v) (c d : R) (i j : ι) (hc : c ≠ 0) + (h : c • v i = d • v j) : i = j := by + have h_single_eq : Finsupp.single i c = Finsupp.single j d := + li <| by simpa [Finsupp.linearCombination_apply] using h + rcases (Finsupp.single_eq_single_iff ..).mp h_single_eq with (⟨H, _⟩ | ⟨hc, _⟩) + · exact H + · contradiction + +section Subtype + +/-! The following lemmas use the subtype defined by a set in `M` as the index set `ι`. -/ + +theorem LinearIndependent.disjoint_span_image (hv : LinearIndependent R v) {s t : Set ι} + (hs : Disjoint s t) : Disjoint (Submodule.span R <| v '' s) (Submodule.span R <| v '' t) := by + simp only [disjoint_def, Finsupp.mem_span_image_iff_linearCombination] + rintro _ ⟨l₁, hl₁, rfl⟩ ⟨l₂, hl₂, H⟩ + rw [hv.injective_linearCombination.eq_iff] at H; subst l₂ + have : l₁ = 0 := Submodule.disjoint_def.mp (Finsupp.disjoint_supported_supported hs) _ hl₁ hl₂ + simp [this] + +theorem LinearIndependent.not_mem_span_image [Nontrivial R] (hv : LinearIndependent R v) {s : Set ι} + {x : ι} (h : x ∉ s) : v x ∉ Submodule.span R (v '' s) := by + have h' : v x ∈ Submodule.span R (v '' {x}) := by + rw [Set.image_singleton] + exact mem_span_singleton_self (v x) + intro w + apply LinearIndependent.ne_zero x hv + refine disjoint_def.1 (hv.disjoint_span_image ?_) (v x) h' w simpa using h +theorem LinearIndependent.linearCombination_ne_of_not_mem_support [Nontrivial R] + (hv : LinearIndependent R v) {x : ι} (f : ι →₀ R) (h : x ∉ f.support) : + Finsupp.linearCombination R v f ≠ v x := by + replace h : x ∉ (f.support : Set ι) := h + have p := hv.not_mem_span_image h + intro w + rw [← w] at p + rw [Finsupp.span_image_eq_map_linearCombination] at p + simp only [not_exists, not_and, mem_map] at p -- Porting note: `mem_map` isn't currently triggered + exact p f (f.mem_supported_support R) rfl + +@[deprecated (since := "2024-08-29")] alias LinearIndependent.total_ne_of_not_mem_support := + LinearIndependent.linearCombination_ne_of_not_mem_support + +end Subtype + +section repr + +/-- Canonical isomorphism between linear combinations and the span of linearly independent vectors. +-/ +@[simps (config := { rhsMd := default }) symm_apply] +def LinearIndependent.linearCombinationEquiv (hv : LinearIndependent R v) : + (ι →₀ R) ≃ₗ[R] span R (range v) := by + refine LinearEquiv.ofBijective (LinearMap.codRestrict (span R (range v)) + (Finsupp.linearCombination R v) ?_) ⟨hv.codRestrict _, ?_⟩ + · simp_rw [← Finsupp.range_linearCombination]; exact fun c ↦ ⟨c, rfl⟩ + rw [← LinearMap.range_eq_top, LinearMap.range_eq_map, LinearMap.map_codRestrict, + ← LinearMap.range_le_iff_comap, range_subtype, Submodule.map_top, + Finsupp.range_linearCombination] + +@[deprecated (since := "2024-08-29")] noncomputable alias LinearIndependent.totalEquiv := + LinearIndependent.linearCombinationEquiv + +-- Porting note: The original theorem generated by `simps` was +-- different from the theorem on Lean 3, and not simp-normal form. +@[simp] +theorem LinearIndependent.linearCombinationEquiv_apply_coe (hv : LinearIndependent R v) + (l : ι →₀ R) : hv.linearCombinationEquiv l = Finsupp.linearCombination R v l := rfl + +@[deprecated (since := "2024-08-29")] alias LinearIndependent.totalEquiv_apply_coe := + LinearIndependent.linearCombinationEquiv_apply_coe +/-- Linear combination representing a vector in the span of linearly independent vectors. + +Given a family of linearly independent vectors, we can represent any vector in their span as +a linear combination of these vectors. These are provided by this linear map. +It is simply one direction of `LinearIndependent.linearCombinationEquiv`. -/ +def LinearIndependent.repr (hv : LinearIndependent R v) : span R (range v) →ₗ[R] ι →₀ R := + hv.linearCombinationEquiv.symm + +variable (hv : LinearIndependent R v) + +@[simp] +theorem LinearIndependent.linearCombination_repr (x) : + Finsupp.linearCombination R v (hv.repr x) = x := + Subtype.ext_iff.1 (LinearEquiv.apply_symm_apply hv.linearCombinationEquiv x) + +@[deprecated (since := "2024-08-29")] alias LinearIndependent.total_repr := + LinearIndependent.linearCombination_repr + +theorem LinearIndependent.linearCombination_comp_repr : + (Finsupp.linearCombination R v).comp hv.repr = Submodule.subtype _ := + LinearMap.ext <| hv.linearCombination_repr + +@[deprecated (since := "2024-08-29")] alias LinearIndependent.total_comp_repr := + LinearIndependent.linearCombination_comp_repr + +theorem LinearIndependent.repr_ker : LinearMap.ker hv.repr = ⊥ := by + rw [LinearIndependent.repr, LinearEquiv.ker] + +theorem LinearIndependent.repr_range : LinearMap.range hv.repr = ⊤ := by + rw [LinearIndependent.repr, LinearEquiv.range] + +theorem LinearIndependent.repr_eq {l : ι →₀ R} {x : span R (range v)} + (eq : Finsupp.linearCombination R v l = ↑x) : hv.repr x = l := by + have : + ↑((LinearIndependent.linearCombinationEquiv hv : (ι →₀ R) →ₗ[R] span R (range v)) l) = + Finsupp.linearCombination R v l := + rfl + have : (LinearIndependent.linearCombinationEquiv hv : (ι →₀ R) →ₗ[R] span R (range v)) l = x := by + rw [eq] at this + exact Subtype.ext_iff.2 this + rw [← LinearEquiv.symm_apply_apply hv.linearCombinationEquiv l] + rw [← this] + rfl + +theorem LinearIndependent.repr_eq_single (i) (x : span R (range v)) (hx : ↑x = v i) : + hv.repr x = Finsupp.single i 1 := by + apply hv.repr_eq + simp [Finsupp.linearCombination_single, hx] + +theorem LinearIndependent.span_repr_eq [Nontrivial R] (x) : + Span.repr R (Set.range v) x = + (hv.repr x).equivMapDomain (Equiv.ofInjective _ hv.injective) := by + have p : + (Span.repr R (Set.range v) x).equivMapDomain (Equiv.ofInjective _ hv.injective).symm = + hv.repr x := by + apply (LinearIndependent.linearCombinationEquiv hv).injective + ext + simp only [LinearIndependent.linearCombinationEquiv_apply_coe, Equiv.self_comp_ofInjective_symm, + LinearIndependent.linearCombination_repr, Finsupp.linearCombination_equivMapDomain, + Span.finsupp_linearCombination_repr] + ext ⟨_, ⟨i, rfl⟩⟩ + simp [← p] + +theorem LinearIndependent.not_smul_mem_span (hv : LinearIndependent R v) (i : ι) (a : R) + (ha : a • v i ∈ span R (v '' (univ \ {i}))) : a = 0 := by + rw [Finsupp.span_image_eq_map_linearCombination, mem_map] at ha + rcases ha with ⟨l, hl, e⟩ + rw [linearIndependent_iffₛ.1 hv l (Finsupp.single i a) (by simp [e])] at hl + by_contra hn + exact (not_mem_of_mem_diff (hl <| by simp [hn])) (mem_singleton _) + +/-- See also `iSupIndep_iff_linearIndependent_of_ne_zero`. -/ +theorem LinearIndependent.iSupIndep_span_singleton (hv : LinearIndependent R v) : + iSupIndep fun i => R ∙ v i := by + refine iSupIndep_def.mp fun i => ?_ + rw [disjoint_iff_inf_le] + intro m hm + simp only [mem_inf, mem_span_singleton, iSup_subtype'] at hm + rw [← span_range_eq_iSup] at hm + obtain ⟨⟨r, rfl⟩, hm⟩ := hm + suffices r = 0 by simp [this] + apply hv.not_smul_mem_span i + convert hm + ext + simp + +@[deprecated (since := "2024-11-24")] +alias LinearIndependent.independent_span_singleton := LinearIndependent.iSupIndep_span_singleton -theorem linearIndependent_empty_type [IsEmpty ι] : LinearIndependent R v := - Function.injective_of_subsingleton _ +end repr -variable (R M) in -theorem linearIndependent_empty : LinearIndependent R (fun x => x : (∅ : Set M) → M) := - linearIndependent_empty_type +section union -/-- A subfamily of a linearly independent family (i.e., a composition with an injective map) is a -linearly independent family. -/ -theorem LinearIndependent.comp (h : LinearIndependent R v) (f : ι' → ι) (hf : Injective f) : - LinearIndependent R (v ∘ f) := by - simpa [Function.comp_def] using Function.Injective.comp h (Finsupp.mapDomain_injective hf) +open LinearMap Finsupp -/-- A set of linearly independent vectors in a module `M` over a semiring `K` is also linearly -independent over a subring `R` of `K`. -The implementation uses minimal assumptions about the relationship between `R`, `K` and `M`. -The version where `K` is an `R`-algebra is `LinearIndependent.restrict_scalars_algebras`. - -/ -theorem LinearIndependent.restrict_scalars [Semiring K] [SMulWithZero R K] [Module K M] - [IsScalarTower R K M] (hinj : Function.Injective fun r : R => r • (1 : K)) - (li : LinearIndependent K v) : LinearIndependent R v := by - intro x y hxy - let f := fun r : R => r • (1 : K) - have := @li (x.mapRange f (by simp [f])) (y.mapRange f (by simp [f])) ?_ - · ext i - exact hinj congr($this i) - simpa [Finsupp.linearCombination, f, Finsupp.sum_mapRange_index] +theorem LinearIndependent.image_subtypeₛ {s : Set M} {f : M →ₗ[R] M'} + (hs : LinearIndependent R (fun x ↦ x : s → M)) + (hf_inj : Set.InjOn f (span R s)) : + LinearIndependent R (fun x ↦ x : f '' s → M') := by + rw [← Subtype.range_coe (s := s)] at hf_inj + refine (hs.map_injOn f hf_inj).to_subtype_range' ?_ + simp [Set.range_comp f] + +theorem linearIndependent_inl_union_inr' {v : ι → M} {v' : ι' → M'} + (hv : LinearIndependent R v) (hv' : LinearIndependent R v') : + LinearIndependent R (Sum.elim (inl R M M' ∘ v) (inr R M M' ∘ v')) := by + have : linearCombination R (Sum.elim (inl R M M' ∘ v) (inr R M M' ∘ v')) = + .prodMap (linearCombination R v) (linearCombination R v') ∘ₗ + (sumFinsuppLEquivProdFinsupp R).toLinearMap := by + ext (_ | _) <;> simp [linearCombination_comapDomain] + rw [LinearIndependent, this] + simpa [LinearMap.coe_prodMap] using ⟨hv, hv'⟩ + +theorem LinearIndependent.inl_union_inr {s : Set M} {t : Set M'} + (hs : LinearIndependent R (fun x => x : s → M)) + (ht : LinearIndependent R (fun x => x : t → M')) : + LinearIndependent R (fun x => x : ↥(inl R M M' '' s ∪ inr R M M' '' t) → M × M') := by + nontriviality R + let e : s ⊕ t ≃ ↥(inl R M M' '' s ∪ inr R M M' '' t) := + .ofBijective (Sum.elim (fun i ↦ ⟨_, .inl ⟨_, i.2, rfl⟩⟩) fun i ↦ ⟨_, .inr ⟨_, i.2, rfl⟩⟩) + ⟨by rintro (_|_) (_|_) eq <;> simp [hs.ne_zero, ht.ne_zero] at eq <;> aesop, + by rintro ⟨_, ⟨_, _, rfl⟩ | ⟨_, _, rfl⟩⟩ <;> aesop⟩ + refine (linearIndependent_equiv' e ?_).mp (linearIndependent_inl_union_inr' hs ht) + ext (_ | _) <;> rfl + +end union + +section Maximal + +universe v w + +/-- +A linearly independent family is maximal if there is no strictly larger linearly independent family. +-/ +@[nolint unusedArguments] +def LinearIndependent.Maximal {ι : Type w} {R : Type u} [Semiring R] {M : Type v} [AddCommMonoid M] + [Module R M] {v : ι → M} (_i : LinearIndependent R v) : Prop := + ∀ (s : Set M) (_i' : LinearIndependent R ((↑) : s → M)) (_h : range v ≤ s), range v = s + +/-- An alternative characterization of a maximal linearly independent family, +quantifying over types (in the same universe as `M`) into which the indexing family injects. +-/ +theorem LinearIndependent.maximal_iff {ι : Type w} {R : Type u} [Semiring R] [Nontrivial R] + {M : Type v} [AddCommMonoid M] [Module R M] {v : ι → M} (i : LinearIndependent R v) : + i.Maximal ↔ + ∀ (κ : Type v) (w : κ → M) (_i' : LinearIndependent R w) (j : ι → κ) (_h : w ∘ j = v), + Surjective j := by + constructor + · rintro p κ w i' j rfl + specialize p (range w) i'.coe_range (range_comp_subset_range _ _) + rw [range_comp, ← image_univ (f := w)] at p + exact range_eq_univ.mp (image_injective.mpr i'.injective p) + · intro p w i' h + specialize + p w ((↑) : w → M) i' (fun i => ⟨v i, range_subset_iff.mp h i⟩) + (by + ext + simp) + have q := congr_arg (fun s => ((↑) : w → M) '' s) p.range_eq + dsimp at q + rw [← image_univ, image_image] at q + simpa using q + +variable (R) + +theorem exists_maximal_independent' (s : ι → M) : + ∃ I : Set ι, + (LinearIndependent R fun x : I => s x) ∧ + ∀ J : Set ι, I ⊆ J → (LinearIndependent R fun x : J => s x) → I = J := by + let indep : Set ι → Prop := fun I => LinearIndependent R (s ∘ (↑) : I → M) + let X := { I : Set ι // indep I } + let r : X → X → Prop := fun I J => I.1 ⊆ J.1 + have key : ∀ c : Set X, IsChain r c → indep (⋃ (I : X) (_ : I ∈ c), I) := by + intro c hc + dsimp [indep] + rw [linearIndependent_comp_subtypeₛ] + intro f hfsupp g hgsupp hsum + rcases eq_empty_or_nonempty c with (rfl | hn) + · rw [show f = 0 by simpa using hfsupp, show g = 0 by simpa using hgsupp] + haveI : IsRefl X r := ⟨fun _ => Set.Subset.refl _⟩ + classical + obtain ⟨I, _I_mem, hI⟩ : ∃ I ∈ c, (f.support ∪ g.support : Set ι) ⊆ I := + f.support.coe_union _ ▸ hc.directedOn.exists_mem_subset_of_finset_subset_biUnion hn <| by + simpa using And.intro hfsupp hgsupp + exact linearIndependent_comp_subtypeₛ.mp I.2 f (subset_union_left.trans hI) + g (subset_union_right.trans hI) hsum + have trans : Transitive r := fun I J K => Set.Subset.trans + obtain ⟨⟨I, hli : indep I⟩, hmax : ∀ a, r ⟨I, hli⟩ a → r a ⟨I, hli⟩⟩ := + exists_maximal_of_chains_bounded + (fun c hc => ⟨⟨⋃ I ∈ c, (I : Set ι), key c hc⟩, fun I => Set.subset_biUnion_of_mem⟩) @trans + exact ⟨I, hli, fun J hsub hli => Set.Subset.antisymm hsub (hmax ⟨J, hli⟩ hsub)⟩ + +end Maximal + +theorem surjective_of_linearIndependent_of_span [Nontrivial R] (hv : LinearIndependent R v) + (f : ι' ↪ ι) (hss : range v ⊆ span R (range (v ∘ f))) : Surjective f := by + intro i + let repr : (span R (range (v ∘ f)) : Type _) → ι' →₀ R := (hv.comp f f.injective).repr + let l := (repr ⟨v i, hss (mem_range_self i)⟩).mapDomain f + have h_total_l : Finsupp.linearCombination R v l = v i := by + dsimp only [l] + rw [Finsupp.linearCombination_mapDomain] + rw [(hv.comp f f.injective).linearCombination_repr] + -- Porting note: `rfl` isn't necessary. + have h_total_eq : Finsupp.linearCombination R v l = Finsupp.linearCombination R v + (Finsupp.single i 1) := by + rw [h_total_l, Finsupp.linearCombination_single, one_smul] + have l_eq : l = _ := hv h_total_eq + dsimp only [l] at l_eq + rw [← Finsupp.embDomain_eq_mapDomain] at l_eq + rcases Finsupp.single_of_embDomain_single (repr ⟨v i, _⟩) f i (1 : R) zero_ne_one.symm l_eq with + ⟨i', hi'⟩ + use i' + exact hi'.2 + +theorem eq_of_linearIndependent_of_span_subtype [Nontrivial R] {s t : Set M} + (hs : LinearIndependent R (fun x => x : s → M)) (h : t ⊆ s) (hst : s ⊆ span R t) : s = t := by + let f : t ↪ s := + ⟨fun x => ⟨x.1, h x.2⟩, fun a b hab => Subtype.coe_injective (Subtype.mk.inj hab)⟩ + have h_surj : Surjective f := by + apply surjective_of_linearIndependent_of_span hs f _ + convert hst <;> simp [f, comp_def] + show s = t + apply Subset.antisymm _ h + intro x hx + rcases h_surj ⟨x, hx⟩ with ⟨y, hy⟩ + convert y.mem + rw [← Subtype.mk.inj hy] end Semiring @@ -185,30 +819,16 @@ theorem linearIndependent_iff : theorem linearIndependent_iff' : LinearIndependent R v ↔ - ∀ s : Finset ι, ∀ g : ι → R, ∑ i ∈ s, g i • v i = 0 → ∀ i ∈ s, g i = 0 := - linearIndependent_iff.trans - ⟨fun hf s g hg i his => - have h := - hf (∑ i ∈ s, Finsupp.single i (g i)) <| by - simpa only [map_sum, Finsupp.linearCombination_single] using hg - calc - g i = (Finsupp.lapply i : (ι →₀ R) →ₗ[R] R) (Finsupp.single i (g i)) := by - { rw [Finsupp.lapply_apply, Finsupp.single_eq_same] } - _ = ∑ j ∈ s, (Finsupp.lapply i : (ι →₀ R) →ₗ[R] R) (Finsupp.single j (g j)) := - Eq.symm <| - Finset.sum_eq_single i - (fun j _hjs hji => by rw [Finsupp.lapply_apply, Finsupp.single_eq_of_ne hji]) - fun hnis => hnis.elim his - _ = (∑ j ∈ s, Finsupp.single j (g j)) i := (map_sum ..).symm - _ = 0 := DFunLike.ext_iff.1 h i, - fun hf _ hl => - Finsupp.ext fun _ => - _root_.by_contradiction fun hni => hni <| hf _ _ hl _ <| Finsupp.mem_support_iff.2 hni⟩ + ∀ s : Finset ι, ∀ g : ι → R, ∑ i ∈ s, g i • v i = 0 → ∀ i ∈ s, g i = 0 := by + rw [linearIndependent_iff'ₛ] + refine ⟨fun h s f ↦ ?_, fun h s f g ↦ ?_⟩ + · convert h s f 0; simp_rw [Pi.zero_apply, zero_smul, Finset.sum_const_zero] + · rw [← sub_eq_zero, ← Finset.sum_sub_distrib] + convert h s (f - g) using 3 <;> simp only [Pi.sub_apply, sub_smul, sub_eq_zero] theorem linearIndependent_iff'' : LinearIndependent R v ↔ - ∀ (s : Finset ι) (g : ι → R), (∀ i ∉ s, g i = 0) → - ∑ i ∈ s, g i • v i = 0 → ∀ i, g i = 0 := by + ∀ (s : Finset ι) (g : ι → R), (∀ i ∉ s, g i = 0) → ∑ i ∈ s, g i • v i = 0 → ∀ i, g i = 0 := by classical exact linearIndependent_iff'.trans ⟨fun H s g hg hv i => if his : i ∈ s then H s g hv i his else hg i his, fun H s g hg i hi => by @@ -243,12 +863,6 @@ theorem Fintype.not_linearIndependent_iff [Fintype ι] : ¬LinearIndependent R v ↔ ∃ g : ι → R, ∑ i, g i • v i = 0 ∧ ∃ i, g i ≠ 0 := by simpa using not_iff_not.2 Fintype.linearIndependent_iff -lemma LinearIndependent.eq_zero_of_pair {x y : M} (h : LinearIndependent R ![x, y]) - {s t : R} (h' : s • x + t • y = 0) : s = 0 ∧ t = 0 := by - replace h := @h (.single 0 s + .single 1 t) 0 ?_ - · exact ⟨by simpa using congr($h 0), by simpa using congr($h 1)⟩ - simpa - /-- Also see `LinearIndependent.pair_iff'` for a simpler version over fields. -/ lemma LinearIndependent.pair_iff {x y : M} : LinearIndependent R ![x, y] ↔ ∀ (s t : R), s • x + t • y = 0 → s = 0 ∧ t = 0 := by @@ -260,50 +874,33 @@ lemma LinearIndependent.pair_iff {x y : M} : fin_cases i exacts [(h _ _ hg).1, (h _ _ hg).2] -/-- A family is linearly independent if and only if all of its finite subfamily is -linearly independent. -/ -theorem linearIndependent_iff_finset_linearIndependent : - LinearIndependent R v ↔ ∀ (s : Finset ι), LinearIndependent R (v ∘ (Subtype.val : s → ι)) := - ⟨fun H _ ↦ H.comp _ Subtype.val_injective, fun H ↦ linearIndependent_iff'.2 fun s g hg i hi ↦ - Fintype.linearIndependent_iff.1 (H s) (g ∘ Subtype.val) - (hg ▸ Finset.sum_attach s fun j ↦ g j • v j) ⟨i, hi⟩⟩ - -theorem LinearIndependent.coe_range (i : LinearIndependent R v) : - LinearIndependent R ((↑) : range v → M) := by simpa using i.comp _ (rangeSplitting_injective v) +/-- If the kernel of a linear map is disjoint from the span of a family of vectors, +then the family is linearly independent iff it is linearly independent after composing with +the linear map. -/ +protected theorem LinearMap.linearIndependent_iff_of_disjoint (f : M →ₗ[R] M') + (hf_inj : Disjoint (span R (Set.range v)) (LinearMap.ker f)) : + LinearIndependent R (f ∘ v) ↔ LinearIndependent R v := + f.linearIndependent_iff_of_injOn <| LinearMap.injOn_of_disjoint_ker le_rfl hf_inj /-- If `v` is a linearly independent family of vectors and the kernel of a linear map `f` is disjoint with the submodule spanned by the vectors of `v`, then `f ∘ v` is a linearly independent family of vectors. See also `LinearIndependent.map'` for a special case assuming `ker f = ⊥`. -/ theorem LinearIndependent.map (hv : LinearIndependent R v) {f : M →ₗ[R] M'} - (hf_inj : Disjoint (span R (range v)) (LinearMap.ker f)) : LinearIndependent R (f ∘ v) := by - rw [disjoint_iff_inf_le, ← Set.image_univ, Finsupp.span_image_eq_map_linearCombination, - map_inf_eq_map_inf_comap, map_le_iff_le_comap, comap_bot, Finsupp.supported_univ, top_inf_eq] - at hf_inj - rw [linearIndependent_iff_ker] at hv ⊢ - rw [hv, le_bot_iff] at hf_inj - rw [Finsupp.linearCombination_linear_comp, LinearMap.ker_comp, hf_inj] - -/-- If `v` is an injective family of vectors such that `f ∘ v` is linearly independent, then `v` - spans a submodule disjoint from the kernel of `f` -/ -theorem Submodule.range_ker_disjoint {f : M →ₗ[R] M'} - (hv : LinearIndependent R (f ∘ v)) : - Disjoint (span R (range v)) (LinearMap.ker f) := by - rw [linearIndependent_iff_ker, Finsupp.linearCombination_linear_comp, LinearMap.ker_comp] at hv - rw [disjoint_iff_inf_le, ← Set.image_univ, Finsupp.span_image_eq_map_linearCombination, - map_inf_eq_map_inf_comap, hv, inf_bot_eq, map_bot] + (hf_inj : Disjoint (span R (range v)) (LinearMap.ker f)) : LinearIndependent R (f ∘ v) := + (f.linearIndependent_iff_of_disjoint hf_inj).mpr hv /-- An injective linear map sends linearly independent families of vectors to linearly independent families of vectors. See also `LinearIndependent.map` for a more general statement. -/ theorem LinearIndependent.map' (hv : LinearIndependent R v) (f : M →ₗ[R] M') (hf_inj : LinearMap.ker f = ⊥) : LinearIndependent R (f ∘ v) := - hv.map <| by simp [hf_inj] + hv.map <| by simp_rw [hf_inj, disjoint_bot_right] /-- If `M / R` and `M' / R'` are modules, `i : R' → R` is a map, `j : M →+ M'` is a monoid map, such that they send non-zero elements to non-zero elements, and compatible with the scalar multiplications on `M` and `M'`, then `j` sends linearly independent families of vectors to linearly independent families of vectors. As a special case, taking `R = R'` it is `LinearIndependent.map'`. -/ -theorem LinearIndependent.map_of_injective_injective {R' : Type*} {M' : Type*} +theorem LinearIndependent.map_of_injective_injective {R' M' : Type*} [Ring R'] [AddCommGroup M'] [Module R' M'] (hv : LinearIndependent R v) (i : R' → R) (j : M →+ M') (hi : ∀ r, i r = 0 → r = 0) (hj : ∀ m, j m = 0 → m = 0) (hc : ∀ (r : R') (m : M), j (i r • m) = r • j m) : LinearIndependent R' (j ∘ v) := by @@ -327,48 +924,11 @@ theorem LinearIndependent.map_of_surjective_injective {R' : Type*} {M' : Type*} rwa [hi', i.map_zero] at h rw [hc (i' r) m, hi'] -/-- If the image of a family of vectors under a linear map is linearly independent, then so is -the original family. -/ -theorem LinearIndependent.of_comp (f : M →ₗ[R] M') (hfv : LinearIndependent R (f ∘ v)) : - LinearIndependent R v := - linearIndependent_iff'.2 fun s g hg i his => - have : (∑ i ∈ s, g i • f (v i)) = 0 := by - simp_rw [← map_smul, ← map_sum, hg, f.map_zero] - linearIndependent_iff'.1 hfv s g this i his - /-- If `f` is an injective linear map, then the family `f ∘ v` is linearly independent if and only if the family `v` is linearly independent. -/ protected theorem LinearMap.linearIndependent_iff (f : M →ₗ[R] M') (hf_inj : LinearMap.ker f = ⊥) : LinearIndependent R (f ∘ v) ↔ LinearIndependent R v := - ⟨fun h => h.of_comp f, fun h => h.map <| by simp only [hf_inj, disjoint_bot_right]⟩ - -@[nontriviality] -theorem linearIndependent_of_subsingleton [Subsingleton R] : LinearIndependent R v := - linearIndependent_iff.2 fun _l _hl => Subsingleton.elim _ _ - -theorem linearIndependent_equiv (e : ι ≃ ι') {f : ι' → M} : - LinearIndependent R (f ∘ e) ↔ LinearIndependent R f := - ⟨fun h => Function.comp_id f ▸ e.self_comp_symm ▸ h.comp _ e.symm.injective, fun h => - h.comp _ e.injective⟩ - -theorem linearIndependent_equiv' (e : ι ≃ ι') {f : ι' → M} {g : ι → M} (h : f ∘ e = g) : - LinearIndependent R g ↔ LinearIndependent R f := - h ▸ linearIndependent_equiv e - -theorem linearIndependent_subtype_range {ι} {f : ι → M} (hf : Injective f) : - LinearIndependent R ((↑) : range f → M) ↔ LinearIndependent R f := - Iff.symm <| linearIndependent_equiv' (Equiv.ofInjective f hf) rfl - -alias ⟨LinearIndependent.of_subtype_range, _⟩ := linearIndependent_subtype_range - -theorem linearIndependent_image {ι} {s : Set ι} {f : ι → M} (hf : Set.InjOn f s) : - (LinearIndependent R fun x : s => f x) ↔ LinearIndependent R fun x : f '' s => (x : M) := - linearIndependent_equiv' (Equiv.Set.imageOfInjOn _ _ hf) rfl - -theorem linearIndependent_span (hs : LinearIndependent R v) : - LinearIndependent R (M := span R (range v)) - (fun i : ι => ⟨v i, subset_span (mem_range_self i)⟩) := - LinearIndependent.of_comp (span R (range v)).subtype hs + f.linearIndependent_iff_of_disjoint <| by simp_rw [hf_inj, disjoint_bot_right] /-- See `LinearIndependent.fin_cons` for a family of elements in a vector space. -/ theorem LinearIndependent.fin_cons' {m : ℕ} (x : M) (v : Fin m → M) (hli : LinearIndependent R v) @@ -383,45 +943,15 @@ theorem LinearIndependent.fin_cons' {m : ℕ} (x : M) (v : Fin m → M) (hli : L rw [this, zero_smul, zero_add] at total_eq exact Fin.cases this (hli _ total_eq) j -/-- Every finite subset of a linearly independent set is linearly independent. -/ -theorem linearIndependent_finset_map_embedding_subtype (s : Set M) - (li : LinearIndependent R ((↑) : s → M)) (t : Finset s) : - LinearIndependent R ((↑) : Finset.map (Embedding.subtype s) t → M) := by - let f : t.map (Embedding.subtype s) → s := fun x => - ⟨x.1, by - obtain ⟨x, h⟩ := x - rw [Finset.mem_map] at h - obtain ⟨a, _ha, rfl⟩ := h - simp only [Subtype.coe_prop, Embedding.coe_subtype]⟩ - convert LinearIndependent.comp li f ?_ - rintro ⟨x, hx⟩ ⟨y, hy⟩ - rw [Finset.mem_map] at hx hy - obtain ⟨a, _ha, rfl⟩ := hx - obtain ⟨b, _hb, rfl⟩ := hy - simp only [f, imp_self, Subtype.mk_eq_mk] - - section Subtype /-! The following lemmas use the subtype defined by a set in `M` as the index set `ι`. -/ - theorem linearIndependent_comp_subtype {s : Set ι} : LinearIndependent R (v ∘ (↑) : s → M) ↔ - ∀ l ∈ Finsupp.supported R R s, (Finsupp.linearCombination R v) l = 0 → l = 0 := by - simp only [linearIndependent_iff, (· ∘ ·), Finsupp.mem_supported, Finsupp.linearCombination_apply, - Set.subset_def, Finset.mem_coe] - constructor - · intro h l hl₁ hl₂ - exact (Finsupp.subtypeDomain_eq_zero_iff hl₁).1 <| - h (l.subtypeDomain s) ((Finsupp.sum_subtypeDomain_index hl₁).trans hl₂) - · intro h l hl - refine Finsupp.embDomain_eq_zero.1 (h (l.embDomain <| Function.Embedding.subtype s) ?_ ?_) - · suffices ∀ i hi, ¬l ⟨i, hi⟩ = 0 → i ∈ s by simpa - intros - assumption - · rwa [Finsupp.embDomain_eq_mapDomain, Finsupp.sum_mapDomain_index] - exacts [fun _ => zero_smul _ _, fun _ _ _ => add_smul _ _ _] + ∀ l ∈ Finsupp.supported R R s, (Finsupp.linearCombination R v) l = 0 → l = 0 := + linearIndependent_comp_subtypeₛ.trans ⟨fun h l hl ↦ h l hl 0 (zero_mem _), fun h f hf g hg eq ↦ + sub_eq_zero.mp (h (f - g) (sub_mem hf hg) <| by rw [map_sub, eq, sub_self])⟩ theorem linearDependent_comp_subtype' {s : Set ι} : ¬LinearIndependent R (v ∘ (↑) : s → M) ↔ @@ -445,60 +975,19 @@ theorem linearIndependent_comp_subtype_disjoint {s : Set ι} : rw [linearIndependent_comp_subtype, LinearMap.disjoint_ker] theorem linearIndependent_subtype_disjoint {s : Set M} : - LinearIndependent R (fun x => x : s → M) ↔ + LinearIndependent R (fun x ↦ x : s → M) ↔ Disjoint (Finsupp.supported R R s) (LinearMap.ker <| Finsupp.linearCombination R id) := by apply linearIndependent_comp_subtype_disjoint (v := id) theorem linearIndependent_iff_linearCombinationOn {s : Set M} : - LinearIndependent R (fun x => x : s → M) ↔ - (LinearMap.ker <| Finsupp.linearCombinationOn M M R id s) = ⊥ := by - rw [Finsupp.linearCombinationOn, LinearMap.ker, LinearMap.comap_codRestrict, Submodule.map_bot, - comap_bot, LinearMap.ker_comp, linearIndependent_subtype_disjoint, disjoint_iff_inf_le, - ← map_comap_subtype, map_le_iff_le_comap, comap_bot, ker_subtype, le_bot_iff] + LinearIndependent R (fun x ↦ x : s → M) ↔ + (LinearMap.ker <| Finsupp.linearCombinationOn M M R id s) = ⊥ := + linearIndependent_iff_linearCombinationOnₛ.trans <| + LinearMap.ker_eq_bot (M := Finsupp.supported R R s).symm @[deprecated (since := "2024-08-29")] alias linearIndependent_iff_totalOn := linearIndependent_iff_linearCombinationOn -theorem LinearIndependent.restrict_of_comp_subtype {s : Set ι} - (hs : LinearIndependent R (v ∘ (↑) : s → M)) : LinearIndependent R (s.restrict v) := - hs - -theorem LinearIndependent.mono {t s : Set M} (h : t ⊆ s) : - LinearIndependent R (fun x => x : s → M) → LinearIndependent R (fun x => x : t → M) := by - simp only [linearIndependent_subtype_disjoint] - exact Disjoint.mono_left (Finsupp.supported_mono h) - -theorem linearIndependent_of_finite (s : Set M) - (H : ∀ t ⊆ s, Set.Finite t → LinearIndependent R (fun x => x : t → M)) : - LinearIndependent R (fun x => x : s → M) := - linearIndependent_subtype.2 fun l hl => - linearIndependent_subtype.1 (H _ hl (Finset.finite_toSet _)) l (Subset.refl _) - -theorem linearIndependent_iUnion_of_directed {η : Type*} {s : η → Set M} (hs : Directed (· ⊆ ·) s) - (h : ∀ i, LinearIndependent R (fun x => x : s i → M)) : - LinearIndependent R (fun x => x : (⋃ i, s i) → M) := by - by_cases hη : Nonempty η - · refine linearIndependent_of_finite (⋃ i, s i) fun t ht ft => ?_ - rcases finite_subset_iUnion ft ht with ⟨I, fi, hI⟩ - rcases hs.finset_le fi.toFinset with ⟨i, hi⟩ - exact (h i).mono (Subset.trans hI <| iUnion₂_subset fun j hj => hi j (fi.mem_toFinset.2 hj)) - · refine (linearIndependent_empty R M).mono (t := iUnion (s ·)) ?_ - rintro _ ⟨_, ⟨i, _⟩, _⟩ - exact hη ⟨i⟩ - -theorem linearIndependent_sUnion_of_directed {s : Set (Set M)} (hs : DirectedOn (· ⊆ ·) s) - (h : ∀ a ∈ s, LinearIndependent R ((↑) : ((a : Set M) : Type _) → M)) : - LinearIndependent R (fun x => x : ⋃₀ s → M) := by - rw [sUnion_eq_iUnion] - exact linearIndependent_iUnion_of_directed hs.directed_val (by simpa using h) - -theorem linearIndependent_biUnion_of_directed {η} {s : Set η} {t : η → Set M} - (hs : DirectedOn (t ⁻¹'o (· ⊆ ·)) s) (h : ∀ a ∈ s, LinearIndependent R (fun x => x : t a → M)) : - LinearIndependent R (fun x => x : (⋃ a ∈ s, t a) → M) := by - rw [biUnion_eq_iUnion] - exact - linearIndependent_iUnion_of_directed (directed_comp.2 <| hs.directed_val) (by simpa using h) - end Subtype end Module @@ -506,82 +995,11 @@ end Module /-! ### Properties which require `Ring R` -/ -section Module - -variable {v : ι → M} -variable [Ring R] [AddCommGroup M] [AddCommGroup M'] -variable [Module R M] [Module R M'] - -@[deprecated (since := "2024-08-29")] alias linearIndependent_iff_injective_total := - linearIndependent_iff_injective_linearCombination - -@[deprecated (since := "2024-08-29")] alias LinearIndependent.injective_total := - LinearIndependent.injective_linearCombination - -theorem LinearIndependent.injective [Nontrivial R] (hv : LinearIndependent R v) : Injective v := by - simpa [Function.comp_def] - using Function.Injective.comp hv (Finsupp.single_left_injective one_ne_zero) - -theorem LinearIndependent.to_subtype_range {ι} {f : ι → M} (hf : LinearIndependent R f) : - LinearIndependent R ((↑) : range f → M) := by - nontriviality R - exact (linearIndependent_subtype_range hf.injective).2 hf - -theorem LinearIndependent.to_subtype_range' {ι} {f : ι → M} (hf : LinearIndependent R f) {t} - (ht : range f = t) : LinearIndependent R ((↑) : t → M) := - ht ▸ hf.to_subtype_range - -theorem LinearIndependent.image_of_comp {ι ι'} (s : Set ι) (f : ι → ι') (g : ι' → M) - (hs : LinearIndependent R fun x : s => g (f x)) : - LinearIndependent R fun x : f '' s => g x := by - nontriviality R - have : InjOn f s := injOn_iff_injective.2 hs.injective.of_comp - exact (linearIndependent_equiv' (Equiv.Set.imageOfInjOn f s this) rfl).1 hs - -theorem LinearIndependent.image {ι} {s : Set ι} {f : ι → M} - (hs : LinearIndependent R fun x : s => f x) : - LinearIndependent R fun x : f '' s => (x : M) := by - convert LinearIndependent.image_of_comp s f id hs - -theorem LinearIndependent.group_smul {G : Type*} [hG : Group G] [DistribMulAction G R] - [DistribMulAction G M] [IsScalarTower G R M] [SMulCommClass G R M] {v : ι → M} - (hv : LinearIndependent R v) (w : ι → G) : LinearIndependent R (w • v) := by - rw [linearIndependent_iff''] at hv ⊢ - intro s g hgs hsum i - refine (smul_eq_zero_iff_eq (w i)).1 ?_ - refine hv s (fun i => w i • g i) (fun i hi => ?_) ?_ i - · dsimp only - exact (hgs i hi).symm ▸ smul_zero _ - · rw [← hsum, Finset.sum_congr rfl _] - intros - dsimp - rw [smul_assoc, smul_comm] - --- This lemma cannot be proved with `LinearIndependent.group_smul` since the action of --- `Rˣ` on `R` is not commutative. -theorem LinearIndependent.units_smul {v : ι → M} (hv : LinearIndependent R v) (w : ι → Rˣ) : - LinearIndependent R (w • v) := by - rw [linearIndependent_iff''] at hv ⊢ - intro s g hgs hsum i - rw [← (w i).mul_left_eq_zero] - refine hv s (fun i => g i • (w i : R)) (fun i hi => ?_) ?_ i - · dsimp only - exact (hgs i hi).symm ▸ zero_smul _ _ - · rw [← hsum, Finset.sum_congr rfl _] - intros - rw [Pi.smul_apply', smul_assoc, Units.smul_def] - -lemma LinearIndependent.eq_of_pair {x y : M} (h : LinearIndependent R ![x, y]) - {s t s' t' : R} (h' : s • x + t • y = s' • x + t' • y) : s = s' ∧ t = t' := by - have : (s - s') • x + (t - t') • y = 0 := by - rw [← sub_eq_zero_of_eq h'] - match_scalars <;> noncomm_ring - simpa [sub_eq_zero] using h.eq_zero_of_pair this +section Module -lemma LinearIndependent.eq_zero_of_pair' {x y : M} (h : LinearIndependent R ![x, y]) - {s t : R} (h' : s • x = t • y) : s = 0 ∧ t = 0 := by - suffices H : s = 0 ∧ 0 = t from ⟨H.1, H.2.symm⟩ - exact h.eq_of_pair (by simpa using h') +variable {v : ι → M} +variable [Ring R] [AddCommGroup M] [AddCommGroup M'] +variable [Module R M] [Module R M'] /-- If two vectors `x` and `y` are linearly independent, so are their linear combinations `a x + b y` and `c x + d y` provided the determinant `a * d - b * c` is nonzero. -/ @@ -600,95 +1018,6 @@ lemma LinearIndependent.linear_combination_pair_of_det_ne_zero {R M : Type*} [Co have J2 : (a * d - b * c) * t = 0 := by linear_combination -b * I1 + a * I2 exact ⟨by simpa [h'] using mul_eq_zero.1 J1, by simpa [h'] using mul_eq_zero.1 J2⟩ -section Maximal - -universe v w - -/-- -A linearly independent family is maximal if there is no strictly larger linearly independent family. --/ -@[nolint unusedArguments] -def LinearIndependent.Maximal {ι : Type w} {R : Type u} [Ring R] {M : Type v} [AddCommGroup M] - [Module R M] {v : ι → M} (_i : LinearIndependent R v) : Prop := - ∀ (s : Set M) (_i' : LinearIndependent R ((↑) : s → M)) (_h : range v ≤ s), range v = s - -/-- An alternative characterization of a maximal linearly independent family, -quantifying over types (in the same universe as `M`) into which the indexing family injects. --/ -theorem LinearIndependent.maximal_iff {ι : Type w} {R : Type u} [Ring R] [Nontrivial R] {M : Type v} - [AddCommGroup M] [Module R M] {v : ι → M} (i : LinearIndependent R v) : - i.Maximal ↔ - ∀ (κ : Type v) (w : κ → M) (_i' : LinearIndependent R w) (j : ι → κ) (_h : w ∘ j = v), - Surjective j := by - constructor - · rintro p κ w i' j rfl - specialize p (range w) i'.coe_range (range_comp_subset_range _ _) - rw [range_comp, ← image_univ (f := w)] at p - exact range_eq_univ.mp (image_injective.mpr i'.injective p) - · intro p w i' h - specialize - p w ((↑) : w → M) i' (fun i => ⟨v i, range_subset_iff.mp h i⟩) - (by - ext - simp) - have q := congr_arg (fun s => ((↑) : w → M) '' s) p.range_eq - dsimp at q - rw [← image_univ, image_image] at q - simpa using q - -end Maximal - -/-- Linear independent families are injective, even if you multiply either side. -/ -theorem LinearIndependent.eq_of_smul_apply_eq_smul_apply {M : Type*} [AddCommGroup M] [Module R M] - {v : ι → M} (li : LinearIndependent R v) (c d : R) (i j : ι) (hc : c ≠ 0) - (h : c • v i = d • v j) : i = j := by - let l : ι →₀ R := Finsupp.single i c - Finsupp.single j d - have h_total : Finsupp.linearCombination R v l = 0 := by - simp_rw [l, LinearMap.map_sub, Finsupp.linearCombination_apply] - simp [h] - have h_single_eq : Finsupp.single i c = Finsupp.single j d := by - rw [linearIndependent_iff] at li - simp [eq_add_of_sub_eq' (li l h_total)] - rcases (Finsupp.single_eq_single_iff ..).mp h_single_eq with (⟨H, _⟩ | ⟨hc, _⟩) - · exact H - · contradiction - -section Subtype - -/-! The following lemmas use the subtype defined by a set in `M` as the index set `ι`. -/ - -theorem LinearIndependent.disjoint_span_image (hv : LinearIndependent R v) {s t : Set ι} - (hs : Disjoint s t) : Disjoint (Submodule.span R <| v '' s) (Submodule.span R <| v '' t) := by - simp only [disjoint_def, Finsupp.mem_span_image_iff_linearCombination] - rintro _ ⟨l₁, hl₁, rfl⟩ ⟨l₂, hl₂, H⟩ - rw [hv.injective_linearCombination.eq_iff] at H; subst l₂ - have : l₁ = 0 := Submodule.disjoint_def.mp (Finsupp.disjoint_supported_supported hs) _ hl₁ hl₂ - simp [this] - -theorem LinearIndependent.not_mem_span_image [Nontrivial R] (hv : LinearIndependent R v) {s : Set ι} - {x : ι} (h : x ∉ s) : v x ∉ Submodule.span R (v '' s) := by - have h' : v x ∈ Submodule.span R (v '' {x}) := by - rw [Set.image_singleton] - exact mem_span_singleton_self (v x) - intro w - apply LinearIndependent.ne_zero x hv - refine disjoint_def.1 (hv.disjoint_span_image ?_) (v x) h' w - simpa using h - -theorem LinearIndependent.linearCombination_ne_of_not_mem_support [Nontrivial R] - (hv : LinearIndependent R v) {x : ι} (f : ι →₀ R) (h : x ∉ f.support) : - Finsupp.linearCombination R v f ≠ v x := by - replace h : x ∉ (f.support : Set ι) := h - have p := hv.not_mem_span_image h - intro w - rw [← w] at p - rw [Finsupp.span_image_eq_map_linearCombination] at p - simp only [not_exists, not_and, mem_map] at p -- Porting note: `mem_map` isn't currently triggered - exact p f (f.mem_supported_support R) rfl - -@[deprecated (since := "2024-08-29")] alias LinearIndependent.total_ne_of_not_mem_support := - LinearIndependent.linearCombination_ne_of_not_mem_support - theorem linearIndependent_sum {v : ι ⊕ ι' → M} : LinearIndependent R v ↔ LinearIndependent R (v ∘ Sum.inl) ∧ @@ -782,110 +1111,11 @@ theorem linearIndependent_iUnion_finite {η : Type*} {ιs : η → Type*} {f : rw [range_sigma_eq_iUnion_range] apply linearIndependent_iUnion_finite_subtype (fun j => (hindep j).to_subtype_range) hd -end Subtype - -section repr - -variable (hv : LinearIndependent R v) - -/-- Canonical isomorphism between linear combinations and the span of linearly independent vectors. --/ -@[simps (config := { rhsMd := default }) symm_apply] -def LinearIndependent.linearCombinationEquiv (hv : LinearIndependent R v) : - (ι →₀ R) ≃ₗ[R] span R (range v) := by - apply LinearEquiv.ofBijective (LinearMap.codRestrict (span R (range v)) - (Finsupp.linearCombination R v) _) - constructor - · rw [← LinearMap.ker_eq_bot, LinearMap.ker_codRestrict, ← linearIndependent_iff_ker] - · apply hv - · intro l - rw [← Finsupp.range_linearCombination] - rw [LinearMap.mem_range] - apply mem_range_self l - · rw [← LinearMap.range_eq_top, LinearMap.range_eq_map, LinearMap.map_codRestrict, - ← LinearMap.range_le_iff_comap, range_subtype, Submodule.map_top] - rw [Finsupp.range_linearCombination] - -@[deprecated (since := "2024-08-29")] noncomputable alias LinearIndependent.totalEquiv := - LinearIndependent.linearCombinationEquiv - --- Porting note: The original theorem generated by `simps` was --- different from the theorem on Lean 3, and not simp-normal form. -@[simp] -theorem LinearIndependent.linearCombinationEquiv_apply_coe (hv : LinearIndependent R v) - (l : ι →₀ R) : hv.linearCombinationEquiv l = Finsupp.linearCombination R v l := rfl - -@[deprecated (since := "2024-08-29")] alias LinearIndependent.totalEquiv_apply_coe := - LinearIndependent.linearCombinationEquiv_apply_coe -/-- Linear combination representing a vector in the span of linearly independent vectors. - -Given a family of linearly independent vectors, we can represent any vector in their span as -a linear combination of these vectors. These are provided by this linear map. -It is simply one direction of `LinearIndependent.linearCombinationEquiv`. -/ -def LinearIndependent.repr (hv : LinearIndependent R v) : span R (range v) →ₗ[R] ι →₀ R := - hv.linearCombinationEquiv.symm - -@[simp] -theorem LinearIndependent.linearCombination_repr (x) : - Finsupp.linearCombination R v (hv.repr x) = x := - Subtype.ext_iff.1 (LinearEquiv.apply_symm_apply hv.linearCombinationEquiv x) - -@[deprecated (since := "2024-08-29")] alias LinearIndependent.total_repr := - LinearIndependent.linearCombination_repr - -theorem LinearIndependent.linearCombination_comp_repr : - (Finsupp.linearCombination R v).comp hv.repr = Submodule.subtype _ := - LinearMap.ext <| hv.linearCombination_repr - -@[deprecated (since := "2024-08-29")] alias LinearIndependent.total_comp_repr := - LinearIndependent.linearCombination_comp_repr - -theorem LinearIndependent.repr_ker : LinearMap.ker hv.repr = ⊥ := by - rw [LinearIndependent.repr, LinearEquiv.ker] - -theorem LinearIndependent.repr_range : LinearMap.range hv.repr = ⊤ := by - rw [LinearIndependent.repr, LinearEquiv.range] - -theorem LinearIndependent.repr_eq {l : ι →₀ R} {x : span R (range v)} - (eq : Finsupp.linearCombination R v l = ↑x) : hv.repr x = l := by - have : - ↑((LinearIndependent.linearCombinationEquiv hv : (ι →₀ R) →ₗ[R] span R (range v)) l) = - Finsupp.linearCombination R v l := - rfl - have : (LinearIndependent.linearCombinationEquiv hv : (ι →₀ R) →ₗ[R] span R (range v)) l = x := by - rw [eq] at this - exact Subtype.ext_iff.2 this - rw [← LinearEquiv.symm_apply_apply hv.linearCombinationEquiv l] - rw [← this] - rfl - -theorem LinearIndependent.repr_eq_single (i) (x : span R (range v)) (hx : ↑x = v i) : - hv.repr x = Finsupp.single i 1 := by - apply hv.repr_eq - simp [Finsupp.linearCombination_single, hx] - -theorem LinearIndependent.span_repr_eq [Nontrivial R] (x) : - Span.repr R (Set.range v) x = - (hv.repr x).equivMapDomain (Equiv.ofInjective _ hv.injective) := by - have p : - (Span.repr R (Set.range v) x).equivMapDomain (Equiv.ofInjective _ hv.injective).symm = - hv.repr x := by - apply (LinearIndependent.linearCombinationEquiv hv).injective - ext - simp only [LinearIndependent.linearCombinationEquiv_apply_coe, Equiv.self_comp_ofInjective_symm, - LinearIndependent.linearCombination_repr, Finsupp.linearCombination_equivMapDomain, - Span.finsupp_linearCombination_repr] - ext ⟨_, ⟨i, rfl⟩⟩ - simp [← p] +open LinearMap theorem linearIndependent_iff_not_smul_mem_span : LinearIndependent R v ↔ ∀ (i : ι) (a : R), a • v i ∈ span R (v '' (univ \ {i})) → a = 0 := - ⟨fun hv i a ha => by - rw [Finsupp.span_image_eq_map_linearCombination, mem_map] at ha - rcases ha with ⟨l, hl, e⟩ - rw [sub_eq_zero.1 (linearIndependent_iff.1 hv (l - Finsupp.single i a) (by simp [e]))] at hl - by_contra hn - exact (not_mem_of_mem_diff (hl <| by simp [hn])) (mem_singleton _), fun H => + ⟨fun hv ↦ hv.not_smul_mem_span, fun H => linearIndependent_iff.2 fun l hl => by ext i; simp only [Finsupp.zero_apply] by_contra hn @@ -899,50 +1129,7 @@ theorem linearIndependent_iff_not_smul_mem_span : simp [hij] · simp [hl]⟩ -/-- See also `iSupIndep_iff_linearIndependent_of_ne_zero`. -/ -theorem LinearIndependent.iSupIndep_span_singleton (hv : LinearIndependent R v) : - iSupIndep fun i => R ∙ v i := by - refine iSupIndep_def.mp fun i => ?_ - rw [disjoint_iff_inf_le] - intro m hm - simp only [mem_inf, mem_span_singleton, iSup_subtype'] at hm - rw [← span_range_eq_iSup] at hm - obtain ⟨⟨r, rfl⟩, hm⟩ := hm - suffices r = 0 by simp [this] - apply linearIndependent_iff_not_smul_mem_span.mp hv i - convert hm - ext - simp - -@[deprecated (since := "2024-11-24")] -alias LinearIndependent.independent_span_singleton := LinearIndependent.iSupIndep_span_singleton - -variable (R) - -theorem exists_maximal_independent' (s : ι → M) : - ∃ I : Set ι, - (LinearIndependent R fun x : I => s x) ∧ - ∀ J : Set ι, I ⊆ J → (LinearIndependent R fun x : J => s x) → I = J := by - let indep : Set ι → Prop := fun I => LinearIndependent R (s ∘ (↑) : I → M) - let X := { I : Set ι // indep I } - let r : X → X → Prop := fun I J => I.1 ⊆ J.1 - have key : ∀ c : Set X, IsChain r c → indep (⋃ (I : X) (_ : I ∈ c), I) := by - intro c hc - dsimp [indep] - rw [linearIndependent_comp_subtype] - intro f hsupport hsum - rcases eq_empty_or_nonempty c with (rfl | hn) - · simpa using hsupport - haveI : IsRefl X r := ⟨fun _ => Set.Subset.refl _⟩ - obtain ⟨I, _I_mem, hI⟩ : ∃ I ∈ c, (f.support : Set ι) ⊆ I := - hc.directedOn.exists_mem_subset_of_finset_subset_biUnion hn hsupport - exact linearIndependent_comp_subtype.mp I.2 f hI hsum - have trans : Transitive r := fun I J K => Set.Subset.trans - obtain ⟨⟨I, hli : indep I⟩, hmax : ∀ a, r ⟨I, hli⟩ a → r a ⟨I, hli⟩⟩ := - exists_maximal_of_chains_bounded - (fun c hc => ⟨⟨⋃ I ∈ c, (I : Set ι), key c hc⟩, fun I => Set.subset_biUnion_of_mem⟩) @trans - exact ⟨I, hli, fun J hsub hli => Set.Subset.antisymm hsub (hmax ⟨J, hli⟩ hsub)⟩ - +variable (R) in theorem exists_maximal_independent (s : ι → M) : ∃ I : Set ι, (LinearIndependent R fun x : I => s x) ∧ @@ -977,68 +1164,11 @@ theorem exists_maximal_independent (s : ι → M) : refine neg_mem (sum_mem fun c hc => smul_mem _ _ (subset_span ⟨c, ?_, rfl⟩)) exact (memJ.mp (supp_f (Finset.erase_subset _ _ hc))).resolve_left (Finset.ne_of_mem_erase hc) -end repr - -theorem surjective_of_linearIndependent_of_span [Nontrivial R] (hv : LinearIndependent R v) - (f : ι' ↪ ι) (hss : range v ⊆ span R (range (v ∘ f))) : Surjective f := by - intro i - let repr : (span R (range (v ∘ f)) : Type _) → ι' →₀ R := (hv.comp f f.injective).repr - let l := (repr ⟨v i, hss (mem_range_self i)⟩).mapDomain f - have h_total_l : Finsupp.linearCombination R v l = v i := by - dsimp only [l] - rw [Finsupp.linearCombination_mapDomain] - rw [(hv.comp f f.injective).linearCombination_repr] - -- Porting note: `rfl` isn't necessary. - have h_total_eq : Finsupp.linearCombination R v l = Finsupp.linearCombination R v - (Finsupp.single i 1) := by - rw [h_total_l, Finsupp.linearCombination_single, one_smul] - have l_eq : l = _ := hv h_total_eq - dsimp only [l] at l_eq - rw [← Finsupp.embDomain_eq_mapDomain] at l_eq - rcases Finsupp.single_of_embDomain_single (repr ⟨v i, _⟩) f i (1 : R) zero_ne_one.symm l_eq with - ⟨i', hi'⟩ - use i' - exact hi'.2 - -theorem eq_of_linearIndependent_of_span_subtype [Nontrivial R] {s t : Set M} - (hs : LinearIndependent R (fun x => x : s → M)) (h : t ⊆ s) (hst : s ⊆ span R t) : s = t := by - let f : t ↪ s := - ⟨fun x => ⟨x.1, h x.2⟩, fun a b hab => Subtype.coe_injective (Subtype.mk.inj hab)⟩ - have h_surj : Surjective f := by - apply surjective_of_linearIndependent_of_span hs f _ - convert hst <;> simp [f, comp_def] - show s = t - apply Subset.antisymm _ h - intro x hx - rcases h_surj ⟨x, hx⟩ with ⟨y, hy⟩ - convert y.mem - rw [← Subtype.mk.inj hy] - -open LinearMap - theorem LinearIndependent.image_subtype {s : Set M} {f : M →ₗ[R] M'} - (hs : LinearIndependent R (fun x => x : s → M)) + (hs : LinearIndependent R (fun x ↦ x : s → M)) (hf_inj : Disjoint (span R s) (LinearMap.ker f)) : - LinearIndependent R (fun x => x : f '' s → M') := by - rw [← Subtype.range_coe (s := s)] at hf_inj - refine (hs.map hf_inj).to_subtype_range' ?_ - simp [Set.range_comp f] - -theorem LinearIndependent.inl_union_inr {s : Set M} {t : Set M'} - (hs : LinearIndependent R (fun x => x : s → M)) - (ht : LinearIndependent R (fun x => x : t → M')) : - LinearIndependent R (fun x => x : ↥(inl R M M' '' s ∪ inr R M M' '' t) → M × M') := by - refine (hs.image_subtype ?_).union (ht.image_subtype ?_) ?_ <;> [simp; simp; skip] - -- Note: https://github.com/leanprover-community/mathlib4/pull/8386 had to change `span_image` into `span_image _` - simp only [span_image _] - simp [disjoint_iff, prod_inf_prod] - -theorem linearIndependent_inl_union_inr' {v : ι → M} {v' : ι' → M'} (hv : LinearIndependent R v) - (hv' : LinearIndependent R v') : - LinearIndependent R (Sum.elim (inl R M M' ∘ v) (inr R M M' ∘ v')) := - (hv.map' (inl R M M') ker_inl).sum_type (hv'.map' (inr R M M') ker_inr) <| by - refine isCompl_range_inl_inr.disjoint.mono ?_ ?_ <;> - simp only [span_le, range_coe, range_comp_subset_range] + LinearIndependent R (fun x ↦ x : f '' s → M') := + hs.image_subtypeₛ <| LinearMap.injOn_of_disjoint_ker le_rfl hf_inj -- See, for example, Keith Conrad's note -- @@ -1256,7 +1386,7 @@ theorem linearIndependent_insert' {ι} {s : Set ι} {a : ι} {f : ι → V} (has linearIndependent_option] -- Porting note: `simp [(· ∘ ·), range_comp f]` → `simp [(· ∘ ·)]; erw [range_comp f ..]; simp` -- https://github.com/leanprover-community/mathlib4/issues/5164 - simp only [Function.comp_def] + simp only [comp_def] erw [range_comp f ((↑) : s → ι)] simp @@ -1357,7 +1487,7 @@ theorem exists_linearIndependent : /-- Indexed version of `exists_linearIndependent`. -/ lemma exists_linearIndependent' (v : ι → V) : - ∃ (κ : Type u') (a : κ → ι), Function.Injective a ∧ + ∃ (κ : Type u') (a : κ → ι), Injective a ∧ Submodule.span K (Set.range (v ∘ a)) = Submodule.span K (Set.range v) ∧ LinearIndependent K (v ∘ a) := by obtain ⟨t, ht, hsp, hli⟩ := exists_linearIndependent K (Set.range v) @@ -1480,3 +1610,5 @@ theorem exists_finite_card_le_of_finite_of_linearIndependent_of_span (ht : t.Fin ⟨this, by rw [← Eq]; exact Finset.card_le_card <| Finset.coe_subset.mp <| by simp [hsu]⟩ end Module + +set_option linter.style.longFile 1700 From ea41783edc9fcf659412cffd3aa0317aa6fde05a Mon Sep 17 00:00:00 2001 From: Nailin Guan <150537269+Thmoas-Guan@users.noreply.github.com> Date: Mon, 6 Jan 2025 09:50:35 +0000 Subject: [PATCH 171/189] chore(FieldTheory/Galois): small cleanup (#20217) Small refinements about Fieldtheory.Galois Mainly about few refine in proof style and some more suggested lemma. --- Mathlib/FieldTheory/Galois/Basic.lean | 7 ++--- Mathlib/FieldTheory/Galois/Infinite.lean | 39 +++++++++++------------- 2 files changed, 20 insertions(+), 26 deletions(-) diff --git a/Mathlib/FieldTheory/Galois/Basic.lean b/Mathlib/FieldTheory/Galois/Basic.lean index 5b30988548f8e..1213496da816d 100644 --- a/Mathlib/FieldTheory/Galois/Basic.lean +++ b/Mathlib/FieldTheory/Galois/Basic.lean @@ -286,11 +286,8 @@ def galoisCoinsertionIntermediateFieldSubgroup [FiniteDimensional F E] [IsGalois GaloisCoinsertion (OrderDual.toDual ∘ (IntermediateField.fixingSubgroup : IntermediateField F E → Subgroup (E ≃ₐ[F] E))) ((IntermediateField.fixedField : Subgroup (E ≃ₐ[F] E) → IntermediateField F E) ∘ - OrderDual.toDual) where - choice H _ := IntermediateField.fixedField H - gc K H := (IntermediateField.le_iff_le H K).symm - u_l_le K := le_of_eq (fixedField_fixingSubgroup K) - choice_eq _ _ := rfl + OrderDual.toDual) := + OrderIso.toGaloisCoinsertion intermediateFieldEquivSubgroup end IsGalois diff --git a/Mathlib/FieldTheory/Galois/Infinite.lean b/Mathlib/FieldTheory/Galois/Infinite.lean index c97aeaad84aa5..d93b6f9fec8dc 100644 --- a/Mathlib/FieldTheory/Galois/Infinite.lean +++ b/Mathlib/FieldTheory/Galois/Infinite.lean @@ -180,19 +180,15 @@ lemma fixingSubgroup_fixedField (H : ClosedSubgroup (K ≃ₐ[k] K)) [IsGalois k /-- The Galois correspondence from intermediate fields to closed subgroups. -/ def IntermediateFieldEquivClosedSubgroup [IsGalois k K] : IntermediateField k K ≃o (ClosedSubgroup (K ≃ₐ[k] K))ᵒᵈ where - toFun := fun L => - { L.fixingSubgroup with - isClosed' := fixingSubgroup_isClosed L } - invFun := fun H => IntermediateField.fixedField H.1 - left_inv := fun L => fixedField_fixingSubgroup L - right_inv := by - intro H + toFun L := ⟨L.fixingSubgroup, fixingSubgroup_isClosed L⟩ + invFun H := IntermediateField.fixedField H.1 + left_inv L := fixedField_fixingSubgroup L + right_inv H := by simp_rw [fixingSubgroup_fixedField H] rfl - map_rel_iff' := by - intro L₁ L₂ - show L₁.fixingSubgroup ≥ L₂.fixingSubgroup ↔ L₁ ≤ L₂ - rw [← fixedField_fixingSubgroup L₂, IntermediateField.le_iff_le, fixedField_fixingSubgroup L₂] + map_rel_iff' {K L} := by + rw [← fixedField_fixingSubgroup L, IntermediateField.le_iff_le, fixedField_fixingSubgroup L] + rfl /-- The Galois correspondence as a `GaloisInsertion` -/ def GaloisInsertionIntermediateFieldClosedSubgroup [IsGalois k K] : @@ -205,16 +201,14 @@ def GaloisInsertionIntermediateFieldClosedSubgroup [IsGalois k K] : /-- The Galois correspondence as a `GaloisCoinsertion` -/ def GaloisCoinsertionIntermediateFieldSubgroup [IsGalois k K] : GaloisCoinsertion (OrderDual.toDual ∘ fun (E : IntermediateField k K) ↦ E.fixingSubgroup) - ((fun (H : Subgroup (K ≃ₐ[k] K)) ↦ IntermediateField.fixedField H) ∘ - OrderDual.toDual) where + ((fun (H : Subgroup (K ≃ₐ[k] K)) ↦ IntermediateField.fixedField H) ∘ OrderDual.toDual) where choice H _ := IntermediateField.fixedField H gc E H := (IntermediateField.le_iff_le H E).symm u_l_le K := le_of_eq (fixedField_fixingSubgroup K) choice_eq _ _ := rfl theorem isOpen_iff_finite (L : IntermediateField k K) [IsGalois k K] : - IsOpen (IntermediateFieldEquivClosedSubgroup L).carrier ↔ - (FiniteDimensional k L) := by + IsOpen L.fixingSubgroup.carrier ↔ FiniteDimensional k L := by refine ⟨fun h ↦ ?_, fun h ↦ IntermediateField.fixingSubgroup_isOpen L⟩ have : (IntermediateFieldEquivClosedSubgroup.toFun L).carrier ∈ nhds 1 := IsOpen.mem_nhds h (congrFun rfl) @@ -234,14 +228,12 @@ theorem isOpen_iff_finite (L : IntermediateField k K) [IsGalois k K] : exact FiniteDimensional.left k L L'.1 theorem normal_iff_isGalois (L : IntermediateField k K) [IsGalois k K] : - Subgroup.Normal (IntermediateFieldEquivClosedSubgroup L).1 ↔ - IsGalois k L := by + L.fixingSubgroup.Normal ↔ IsGalois k L := by refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ · let f : L → IntermediateField k K := fun x => IntermediateField.lift <| - IntermediateField.fixedField <| Subgroup.map (restrictNormalHom - (adjoin k {x.1})) L.fixingSubgroup - have h' (x : K) : (Subgroup.map (restrictNormalHom - (adjoin k {x})) L.fixingSubgroup).Normal := + IntermediateField.fixedField <| Subgroup.map (restrictNormalHom (adjoin k {x.1})) + L.fixingSubgroup + have h' (x : K) : (Subgroup.map (restrictNormalHom (adjoin k {x})) L.fixingSubgroup).Normal := Subgroup.Normal.map h (restrictNormalHom (adjoin k {x})) (restrictNormalHom_surjective K) have n' (l : L) : IsGalois k (IntermediateField.fixedField <| Subgroup.map (restrictNormalHom (adjoin k {l.1})) L.fixingSubgroup) := by @@ -267,4 +259,9 @@ theorem normal_iff_isGalois (L : IntermediateField k K) [IsGalois k K] : · simpa only [IntermediateFieldEquivClosedSubgroup, RelIso.coe_fn_mk, Equiv.coe_fn_mk, ← L.restrictNormalHom_ker] using MonoidHom.normal_ker (restrictNormalHom L) +theorem isOpen_and_normal_iff_finite_and_isGalois (L : IntermediateField k K) [IsGalois k K] : + IsOpen L.fixingSubgroup.carrier ∧ L.fixingSubgroup.Normal ↔ + FiniteDimensional k L ∧ IsGalois k L := by + rw [isOpen_iff_finite, normal_iff_isGalois] + end InfiniteGalois From 886b34020dd14621e8265d3aa4102c9b1fc6a514 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 6 Jan 2025 10:15:37 +0000 Subject: [PATCH 172/189] chore: remove >9 month old deprecations (#20505) --- Mathlib/Algebra/Algebra/Operations.lean | 13 ----- Mathlib/Algebra/Associated/Basic.lean | 2 - Mathlib/Algebra/CharP/Reduced.lean | 3 -- Mathlib/Algebra/GCDMonoid/Basic.lean | 3 -- Mathlib/Algebra/Group/Basic.lean | 18 ------- Mathlib/Algebra/Group/Defs.lean | 2 - Mathlib/Algebra/Group/UniqueProds/Basic.lean | 17 ------- Mathlib/Algebra/Group/Units/Basic.lean | 2 - .../Algebra/GroupWithZero/Units/Basic.lean | 10 ---- .../Algebra/Module/Submodule/Pointwise.lean | 5 -- .../Order/BigOperators/Ring/Finset.lean | 2 - Mathlib/Algebra/Order/Field/Basic.lean | 9 ---- Mathlib/Algebra/Order/Field/Power.lean | 2 - Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean | 3 -- Mathlib/Algebra/Polynomial/Div.lean | 1 - Mathlib/Algebra/Squarefree/Basic.lean | 3 -- Mathlib/AlgebraicGeometry/OpenImmersion.lean | 3 -- .../ProjectiveSpectrum/Scheme.lean | 1 - Mathlib/Analysis/CStarAlgebra/Matrix.lean | 22 -------- .../Analysis/CStarAlgebra/Unitization.lean | 4 -- .../Calculus/ContDiff/FTaylorSeries.lean | 4 -- .../InverseFunctionTheorem/Deriv.lean | 2 - .../InverseFunctionTheorem/FDeriv.lean | 2 - Mathlib/Analysis/Calculus/MeanValue.lean | 9 ---- .../Analysis/Complex/Polynomial/Basic.lean | 3 -- .../Analysis/Fourier/FourierTransform.lean | 8 --- .../InnerProductSpace/LaxMilgram.lean | 1 - .../Analysis/InnerProductSpace/TwoDim.lean | 4 -- Mathlib/Analysis/LocallyConvex/Basic.lean | 11 ---- Mathlib/Analysis/Matrix.lean | 33 ------------ Mathlib/Analysis/Normed/Lp/LpEquiv.lean | 4 -- .../Normed/Module/FiniteDimension.lean | 17 ------- .../NormedSpace/Multilinear/Basic.lean | 50 ------------------- .../NormedSpace/OperatorNorm/Basic.lean | 25 ---------- .../NormedSpace/OperatorNorm/Bilinear.lean | 5 -- .../OperatorNorm/Completeness.lean | 1 - .../NormedSpace/OperatorNorm/Mul.lean | 17 ------- .../NormedSpace/OperatorNorm/NNNorm.lean | 24 --------- .../NormedSpace/OperatorNorm/NormedSpace.lean | 4 -- .../NormedSpace/OperatorNorm/Prod.lean | 2 - Mathlib/Analysis/NormedSpace/RCLike.lean | 4 -- .../Gaussian/FourierTransform.lean | 9 ---- Mathlib/CategoryTheory/Filtered/Basic.lean | 3 -- Mathlib/CategoryTheory/IsConnected.lean | 4 -- .../Limits/FullSubcategory.lean | 6 --- Mathlib/CategoryTheory/Limits/Opposites.lean | 4 -- .../Combinatorics/SimpleGraph/Coloring.lean | 4 -- Mathlib/Data/Complex/Abs.lean | 3 -- Mathlib/Data/Countable/Small.lean | 6 --- Mathlib/Data/Fin/Basic.lean | 6 --- Mathlib/Data/Finset/Card.lean | 5 -- Mathlib/Data/Finset/Empty.lean | 1 - Mathlib/Data/Finset/Lattice/Basic.lean | 4 -- Mathlib/Data/Int/Defs.lean | 3 -- Mathlib/Data/List/Basic.lean | 6 --- Mathlib/Data/List/Defs.lean | 2 - Mathlib/Data/List/Infix.lean | 1 - Mathlib/Data/List/Pairwise.lean | 9 ---- Mathlib/Data/List/Sublists.lean | 5 -- Mathlib/Data/NNReal/Defs.lean | 1 - Mathlib/Data/Nat/Cast/Order/Basic.lean | 3 -- Mathlib/Data/Nat/Digits.lean | 3 -- Mathlib/Data/Rat/Cast/Defs.lean | 2 - Mathlib/Data/Rat/Defs.lean | 16 ------ Mathlib/Data/Real/Sqrt.lean | 5 -- Mathlib/Data/Set/Basic.lean | 7 --- Mathlib/Data/Set/Image.lean | 21 -------- Mathlib/Data/ZMod/Quotient.lean | 3 -- Mathlib/Deprecated/Cardinal/PartENat.lean | 4 -- Mathlib/Deprecated/Submonoid.lean | 1 - Mathlib/FieldTheory/PrimitiveElement.lean | 6 --- .../Euclidean/Angle/Unoriented/Affine.lean | 3 -- .../Manifold/SmoothManifoldWithCorners.lean | 1 - Mathlib/GroupTheory/OrderOfElement.lean | 8 --- .../GroupTheory/SpecificGroups/Cyclic.lean | 16 ------ Mathlib/LinearAlgebra/FiniteDimensional.lean | 4 -- .../LinearAlgebra/FiniteDimensional/Defs.lean | 9 ---- Mathlib/Logic/Basic.lean | 11 ---- .../Constructions/BorelSpace/Real.lean | 15 ------ Mathlib/MeasureTheory/Function/L1Space.lean | 4 -- .../MeasureTheory/Function/SimpleFunc.lean | 1 - .../Function/StronglyMeasurable/Basic.lean | 9 ---- .../MeasureTheory/Integral/PeakFunction.lean | 12 ----- Mathlib/MeasureTheory/Integral/SetToL1.lean | 3 -- .../MeasureTheory/Measure/Haar/Unique.lean | 15 ------ Mathlib/NumberTheory/LSeries/Dirichlet.lean | 3 -- .../Transcendental/Liouville/Residual.lean | 1 - Mathlib/RingTheory/Flat/Basic.lean | 6 --- .../RingTheory/Ideal/Quotient/Operations.lean | 5 -- .../RingTheory/MvPolynomial/Homogeneous.lean | 3 -- Mathlib/RingTheory/PowerBasis.lean | 1 - Mathlib/RingTheory/PrincipalIdealDomain.lean | 6 --- Mathlib/RingTheory/SimpleModule.lean | 6 --- .../UniqueFactorizationDomain/Defs.lean | 5 -- .../UniqueFactorizationDomain/FactorSet.lean | 1 - .../Multiplicity.lean | 5 -- Mathlib/SetTheory/Cardinal/Arithmetic.lean | 6 --- Mathlib/SetTheory/Cardinal/ToNat.lean | 19 ------- .../Topology/Algebra/InfiniteSum/NatInt.lean | 10 ---- Mathlib/Topology/Algebra/UniformRing.lean | 10 ---- Mathlib/Topology/Bases.lean | 6 --- Mathlib/Topology/CompactOpen.lean | 15 ------ Mathlib/Topology/Compactness/Compact.lean | 16 ------ Mathlib/Topology/Constructions.lean | 2 - .../Topology/ContinuousMap/Bounded/Basic.lean | 2 - .../Topology/ContinuousMap/ZeroAtInfty.lean | 1 - Mathlib/Topology/GDelta/Basic.lean | 9 ---- Mathlib/Topology/GDelta/UniformSpace.lean | 1 - Mathlib/Topology/Instances/Discrete.lean | 5 -- Mathlib/Topology/Instances/Int.lean | 1 - Mathlib/Topology/Instances/Irrational.lean | 1 - Mathlib/Topology/Instances/Real.lean | 3 -- Mathlib/Topology/Maps/Basic.lean | 1 - Mathlib/Topology/MetricSpace/Completion.lean | 4 -- Mathlib/Topology/MetricSpace/Polish.lean | 3 -- Mathlib/Topology/MetricSpace/Sequences.lean | 6 --- Mathlib/Topology/Order.lean | 11 ---- Mathlib/Topology/Order/Basic.lean | 10 ---- Mathlib/Topology/Order/OrderClosed.lean | 8 --- Mathlib/Topology/Semicontinuous.lean | 6 --- Mathlib/Topology/Separation/GDelta.lean | 2 - Mathlib/Topology/Separation/Hausdorff.lean | 3 -- Mathlib/Topology/Separation/Regular.lean | 9 ---- Mathlib/Topology/UniformSpace/Basic.lean | 14 ------ Mathlib/Topology/UnitInterval.lean | 2 - 125 files changed, 847 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index 01b7c2e0ecb3f..7a693780a741f 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -144,10 +144,6 @@ theorem smul_mono_left (h : I ≤ J) : I • N ≤ J • N := instance : CovariantClass (Submodule R A) (Submodule R M) HSMul.hSMul LE.le := ⟨fun _ _ => smul_mono le_rfl⟩ -@[deprecated smul_mono_right (since := "2024-03-31")] -protected theorem smul_mono_right (h : N ≤ P) : I • N ≤ I • P := - _root_.smul_mono_right I h - variable (I J N P) @[simp] @@ -180,10 +176,6 @@ protected theorem smul_assoc {B} [Semiring B] [Module R B] [Module A B] [Module (fun j hj n hn ↦ (smul_assoc r j n).symm ▸ smul_mem_smul (smul_mem_smul hr hj) hn) fun m₁ m₂ ↦ (smul_add r m₁ m₂) ▸ add_mem) -@[deprecated smul_inf_le (since := "2024-03-31")] -protected theorem smul_inf_le (M₁ M₂ : Submodule R M) : - I • (M₁ ⊓ M₂) ≤ I • M₁ ⊓ I • M₂ := smul_inf_le _ _ _ - theorem smul_iSup {ι : Sort*} {I : Submodule R A} {t : ι → Submodule R M} : I • (⨆ i, t i)= ⨆ i, I • t i := toAddSubmonoid_injective <| by @@ -196,11 +188,6 @@ theorem iSup_smul {ι : Sort*} {t : ι → Submodule R A} {N : Submodule R M} : (by simp_rw [zero_smul]; apply zero_mem) fun x y ↦ by simp_rw [add_smul]; apply add_mem) (iSup_le fun i ↦ Submodule.smul_mono_left <| le_iSup _ i) -@[deprecated smul_iInf_le (since := "2024-03-31")] -protected theorem smul_iInf_le {ι : Sort*} {I : Submodule R A} {t : ι → Submodule R M} : - I • iInf t ≤ ⨅ i, I • t i := - smul_iInf_le - protected theorem one_smul : (1 : Submodule R A) • N = N := by refine le_antisymm (smul_le.mpr fun r hr m hm ↦ ?_) fun m hm ↦ ?_ · obtain ⟨r, rfl⟩ := hr diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index 4484964fb4604..12ae9d5fdd160 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -581,7 +581,6 @@ theorem mk_le_mk_of_dvd {a b : M} : a ∣ b → Associates.mk a ≤ Associates.m theorem mk_le_mk_iff_dvd {a b : M} : Associates.mk a ≤ Associates.mk b ↔ a ∣ b := mk_dvd_mk -@[deprecated (since := "2024-03-16")] alias mk_le_mk_iff_dvd_iff := mk_le_mk_iff_dvd @[simp] theorem isPrimal_mk {a : M} : IsPrimal (Associates.mk a) ↔ IsPrimal a := by @@ -591,7 +590,6 @@ theorem isPrimal_mk {a : M} : IsPrimal (Associates.mk a) ↔ IsPrimal a := by exact ⟨a₁, a₂ * u, h₁, Units.mul_right_dvd.mpr h₂, mul_assoc _ _ _⟩ · exact ⟨a₁, a₂, h₁, h₂, congr_arg _ eq⟩ -@[deprecated (since := "2024-03-16")] alias isPrimal_iff := isPrimal_mk @[simp] theorem decompositionMonoid_iff : DecompositionMonoid (Associates M) ↔ DecompositionMonoid M := by diff --git a/Mathlib/Algebra/CharP/Reduced.lean b/Mathlib/Algebra/CharP/Reduced.lean index deea5f7b1eb53..90e5aa6b3aa17 100644 --- a/Mathlib/Algebra/CharP/Reduced.lean +++ b/Mathlib/Algebra/CharP/Reduced.lean @@ -44,6 +44,3 @@ theorem ExpChar.pow_prime_pow_mul_eq_one_iff (p k m : ℕ) [ExpChar R p] (x : R) rw [pow_mul'] convert ← (iterateFrobenius_inj R p k).eq_iff apply map_one - -@[deprecated (since := "2024-02-02")] -alias CharP.pow_prime_pow_mul_eq_one_iff := ExpChar.pow_prime_pow_mul_eq_one_iff diff --git a/Mathlib/Algebra/GCDMonoid/Basic.lean b/Mathlib/Algebra/GCDMonoid/Basic.lean index 16298274aed73..57c84854496c4 100644 --- a/Mathlib/Algebra/GCDMonoid/Basic.lean +++ b/Mathlib/Algebra/GCDMonoid/Basic.lean @@ -805,9 +805,6 @@ theorem lcm_eq_of_associated_right [NormalizedGCDMonoid α] {m n : α} (h : Asso end LCM -@[deprecated (since := "2024-02-12")] alias GCDMonoid.prime_of_irreducible := Irreducible.prime -@[deprecated (since := "2024-02-12")] alias GCDMonoid.irreducible_iff_prime := irreducible_iff_prime - end GCDMonoid section UniqueUnit diff --git a/Mathlib/Algebra/Group/Basic.lean b/Mathlib/Algebra/Group/Basic.lean index 1df564f491efd..bf22f85392085 100644 --- a/Mathlib/Algebra/Group/Basic.lean +++ b/Mathlib/Algebra/Group/Basic.lean @@ -797,14 +797,6 @@ theorem leftInverse_inv_mul_mul_right (c : G) : @[to_additive (attr := simp) natAbs_nsmul_eq_zero] lemma pow_natAbs_eq_one : a ^ n.natAbs = 1 ↔ a ^ n = 1 := by cases n <;> simp -set_option linter.existingAttributeWarning false in -@[to_additive, deprecated pow_natAbs_eq_one (since := "2024-02-14")] -lemma exists_pow_eq_one_of_zpow_eq_one (hn : n ≠ 0) (h : a ^ n = 1) : - ∃ n : ℕ, 0 < n ∧ a ^ n = 1 := ⟨_, Int.natAbs_pos.2 hn, pow_natAbs_eq_one.2 h⟩ - -attribute [deprecated natAbs_nsmul_eq_zero (since := "2024-02-14")] -exists_nsmul_eq_zero_of_zsmul_eq_zero - @[to_additive sub_nsmul] lemma pow_sub (a : G) {m n : ℕ} (h : n ≤ m) : a ^ (m - n) = a ^ m * (a ^ n)⁻¹ := eq_mul_inv_of_mul_eq <| by rw [← pow_add, Nat.sub_add_cancel h] @@ -1041,15 +1033,5 @@ theorem multiplicative_of_isTotal (p : α → Prop) (hswap : ∀ {a b}, p a → end multiplicative -@[deprecated (since := "2024-03-20")] alias div_mul_cancel' := div_mul_cancel -@[deprecated (since := "2024-03-20")] alias mul_div_cancel'' := mul_div_cancel_right -- The name `add_sub_cancel` was reused -- @[deprecated (since := "2024-03-20")] alias add_sub_cancel := add_sub_cancel_right -@[deprecated (since := "2024-03-20")] alias div_mul_cancel''' := div_mul_cancel_right -@[deprecated (since := "2024-03-20")] alias sub_add_cancel'' := sub_add_cancel_right -@[deprecated (since := "2024-03-20")] alias mul_div_cancel''' := mul_div_cancel_left -@[deprecated (since := "2024-03-20")] alias add_sub_cancel' := add_sub_cancel_left -@[deprecated (since := "2024-03-20")] alias mul_div_cancel'_right := mul_div_cancel -@[deprecated (since := "2024-03-20")] alias add_sub_cancel'_right := add_sub_cancel -@[deprecated (since := "2024-03-20")] alias div_mul_cancel'' := div_mul_cancel_left -@[deprecated (since := "2024-03-20")] alias sub_add_cancel' := sub_add_cancel_left diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index 324b405f14be4..9808a37d910f3 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -886,8 +886,6 @@ theorem zpow_natCast (a : G) : ∀ n : ℕ, a ^ (n : ℤ) = a ^ n _ = a ^ n * a := congrArg (· * a) (zpow_natCast a n) _ = a ^ (n + 1) := (pow_succ _ _).symm -@[deprecated (since := "2024-03-20")] alias zpow_coe_nat := zpow_natCast -@[deprecated (since := "2024-03-20")] alias coe_nat_zsmul := natCast_zsmul @[to_additive ofNat_zsmul] lemma zpow_ofNat (a : G) (n : ℕ) : a ^ (ofNat(n) : ℤ) = a ^ OfNat.ofNat n := diff --git a/Mathlib/Algebra/Group/UniqueProds/Basic.lean b/Mathlib/Algebra/Group/UniqueProds/Basic.lean index 41d9de70596fd..2e38b8dc940aa 100644 --- a/Mathlib/Algebra/Group/UniqueProds/Basic.lean +++ b/Mathlib/Algebra/Group/UniqueProds/Basic.lean @@ -625,23 +625,6 @@ instance (priority := 100) of_covariant_left [IsLeftCancelMul G] end TwoUniqueProds -@[deprecated (since := "2024-02-04")] -alias UniqueProds.mulHom_image_of_injective := UniqueProds.of_injective_mulHom -@[deprecated (since := "2024-02-04")] -alias UniqueSums.addHom_image_of_injective := UniqueSums.of_injective_addHom -@[deprecated (since := "2024-02-04")] -alias UniqueProds.mulHom_image_iff := MulEquiv.uniqueProds_iff -@[deprecated (since := "2024-02-04")] -alias UniqueSums.addHom_image_iff := AddEquiv.uniqueSums_iff -@[deprecated (since := "2024-02-04")] -alias TwoUniqueProds.mulHom_image_of_injective := TwoUniqueProds.of_injective_mulHom -@[deprecated (since := "2024-02-04")] -alias TwoUniqueSums.addHom_image_of_injective := TwoUniqueSums.of_injective_addHom -@[deprecated (since := "2024-02-04")] -alias TwoUniqueProds.mulHom_image_iff := MulEquiv.twoUniqueProds_iff -@[deprecated (since := "2024-02-04")] -alias TwoUniqueSums.addHom_image_iff := AddEquiv.twoUniqueSums_iff - instance {ι} (G : ι → Type*) [∀ i, AddZeroClass (G i)] [∀ i, TwoUniqueSums (G i)] : TwoUniqueSums (Π₀ i, G i) := TwoUniqueSums.of_injective_addHom diff --git a/Mathlib/Algebra/Group/Units/Basic.lean b/Mathlib/Algebra/Group/Units/Basic.lean index 3630673d90f59..3cb0df004f9d7 100644 --- a/Mathlib/Algebra/Group/Units/Basic.lean +++ b/Mathlib/Algebra/Group/Units/Basic.lean @@ -474,5 +474,3 @@ attribute [deprecated sub_add_cancel_left (since := "2024-03-20")] IsAddUnit.sub -- @[deprecated (since := "2024-03-20")] alias IsUnit.mul_div_cancel := IsUnit.mul_div_cancel_right -- @[deprecated (since := "2024-03-20")] -- alias IsAddUnit.add_sub_cancel := IsAddUnit.add_sub_cancel_right -@[deprecated (since := "2024-03-20")] alias IsUnit.mul_div_cancel' := IsUnit.mul_div_cancel -@[deprecated (since := "2024-03-20")] alias IsAddUnit.add_sub_cancel' := IsAddUnit.add_sub_cancel diff --git a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean index 7702e037e5cb7..5b98cfd3fad68 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean @@ -315,10 +315,6 @@ lemma div_eq_one_iff_eq (hb : b ≠ 0) : a / b = 1 ↔ a = b := hb.isUnit.div_eq lemma div_mul_cancel_right₀ (hb : b ≠ 0) (a : G₀) : b / (a * b) = a⁻¹ := hb.isUnit.div_mul_cancel_right _ -set_option linter.deprecated false in -@[deprecated div_mul_cancel_right₀ (since := "2024-03-20")] -lemma div_mul_left (hb : b ≠ 0) : b / (a * b) = 1 / a := hb.isUnit.div_mul_left - lemma mul_div_mul_right (a b : G₀) (hc : c ≠ 0) : a * c / (b * c) = a / b := hc.isUnit.mul_div_mul_right _ _ @@ -415,10 +411,6 @@ instance (priority := 100) CommGroupWithZero.toDivisionCommMonoid : lemma div_mul_cancel_left₀ (ha : a ≠ 0) (b : G₀) : a / (a * b) = b⁻¹ := ha.isUnit.div_mul_cancel_left _ -@[deprecated div_mul_cancel_left₀ (since := "2024-03-22")] -lemma div_mul_right (b : G₀) (ha : a ≠ 0) : a / (a * b) = 1 / b := by - simp [div_mul_cancel_left₀ ha] - lemma mul_div_cancel_left_of_imp (h : a = 0 → b = 0) : a * b / a = b := by rw [mul_comm, mul_div_cancel_of_imp h] @@ -485,5 +477,3 @@ noncomputable def commGroupWithZeroOfIsUnitOrEqZero [hM : CommMonoidWithZero M] { groupWithZeroOfIsUnitOrEqZero h, hM with } end NoncomputableDefs - -@[deprecated (since := "2024-03-20")] alias mul_div_cancel' := mul_div_cancel₀ diff --git a/Mathlib/Algebra/Module/Submodule/Pointwise.lean b/Mathlib/Algebra/Module/Submodule/Pointwise.lean index b3c2fa905b8d3..811814d19e270 100644 --- a/Mathlib/Algebra/Module/Submodule/Pointwise.lean +++ b/Mathlib/Algebra/Module/Submodule/Pointwise.lean @@ -347,11 +347,6 @@ instance : CovariantClass (Set S) (Submodule R M) HSMul.hSMul LE.le := ⟨fun _ _ _ le => set_smul_le _ _ _ fun _ _ hr hm => mem_set_smul_of_mem_mem (mem1 := hr) (mem2 := le hm)⟩ -@[deprecated smul_mono_right (since := "2024-03-31")] -theorem set_smul_mono_right {p q : Submodule R M} (le : p ≤ q) : - s • p ≤ s • q := - smul_mono_right s le - lemma set_smul_mono_left {s t : Set S} (le : s ≤ t) : s • N ≤ t • N := set_smul_le _ _ _ fun _ _ hr hm => mem_set_smul_of_mem_mem (mem1 := le hr) diff --git a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean index 607f6e70c1a6c..96a7184d1eef9 100644 --- a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean @@ -217,8 +217,6 @@ lemma IsAbsoluteValue.abv_sum [Semiring R] [OrderedSemiring S] (abv : R → S) [ (f : ι → R) (s : Finset ι) : abv (∑ i ∈ s, f i) ≤ ∑ i ∈ s, abv (f i) := (IsAbsoluteValue.toAbsoluteValue abv).sum_le _ _ -@[deprecated (since := "2024-02-14")] alias abv_sum_le_sum_abv := IsAbsoluteValue.abv_sum - nonrec lemma AbsoluteValue.map_prod [CommSemiring R] [Nontrivial R] [LinearOrderedCommRing S] (abv : AbsoluteValue R S) (f : ι → R) (s : Finset ι) : abv (∏ i ∈ s, f i) = ∏ i ∈ s, abv (f i) := diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index e65d0da3e42fb..d2b9446b43e01 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -166,15 +166,6 @@ theorem one_le_inv_iff : 1 ≤ a⁻¹ ↔ 0 < a ∧ a ≤ 1 := one_le_inv_iff₀ ### Relating two divisions. -/ -@[deprecated (since := "2024-02-16")] alias div_le_div_of_le_of_nonneg := div_le_div_of_nonneg_right -@[deprecated (since := "2024-02-16")] alias div_lt_div_of_lt := div_lt_div_of_pos_right -@[deprecated (since := "2024-02-16")] alias div_le_div_of_le_left := div_le_div_of_nonneg_left -@[deprecated (since := "2024-02-16")] alias div_lt_div_of_lt_left := div_lt_div_of_pos_left - -@[deprecated div_le_div_of_nonneg_right (since := "2024-02-16")] -lemma div_le_div_of_le (hc : 0 ≤ c) (hab : a ≤ b) : a / c ≤ b / c := - div_le_div_of_nonneg_right hab hc - @[deprecated div_le_div_iff_of_pos_right (since := "2024-11-12")] theorem div_le_div_right (hc : 0 < c) : a / c ≤ b / c ↔ a ≤ b := div_le_div_iff_of_pos_right hc diff --git a/Mathlib/Algebra/Order/Field/Power.lean b/Mathlib/Algebra/Order/Field/Power.lean index 0da30f4498bde..ac82917ae0852 100644 --- a/Mathlib/Algebra/Order/Field/Power.lean +++ b/Mathlib/Algebra/Order/Field/Power.lean @@ -55,8 +55,6 @@ theorem zpow_strictAnti (h₀ : 0 < a) (h₁ : a < 1) : StrictAnti (a ^ · : ℤ theorem zpow_lt_iff_lt (hx : 1 < a) : a ^ m < a ^ n ↔ m < n := zpow_lt_zpow_iff_right₀ hx -@[deprecated (since := "2024-02-10")] alias ⟨_, zpow_lt_of_lt⟩ := zpow_lt_iff_lt - @[deprecated zpow_le_zpow_iff_right₀ (since := "2024-10-08")] theorem zpow_le_iff_le (hx : 1 < a) : a ^ m ≤ a ^ n ↔ m ≤ n := zpow_le_zpow_iff_right₀ hx diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean index 21b69f6811d62..bd7669e3210be 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean @@ -197,9 +197,6 @@ instance : AddLeftMono ℚ where @[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] diff --git a/Mathlib/Algebra/Polynomial/Div.lean b/Mathlib/Algebra/Polynomial/Div.lean index 3bcbfbaac6a50..48d16d1d4a6f2 100644 --- a/Mathlib/Algebra/Polynomial/Div.lean +++ b/Mathlib/Algebra/Polynomial/Div.lean @@ -396,7 +396,6 @@ theorem modByMonic_eq_zero_iff_dvd (hq : Monic q) : p %ₘ q = 0 ↔ q ∣ p := degree_eq_natDegree (mt leadingCoeff_eq_zero.2 hrpq0)] at this exact not_lt_of_ge (Nat.le_add_right _ _) (WithBot.coe_lt_coe.1 this)⟩ -@[deprecated (since := "2024-03-23")] alias dvd_iff_modByMonic_eq_zero := modByMonic_eq_zero_iff_dvd /-- See `Polynomial.mul_left_modByMonic` for the other multiplication order. That version, unlike this one, requires commutativity. -/ diff --git a/Mathlib/Algebra/Squarefree/Basic.lean b/Mathlib/Algebra/Squarefree/Basic.lean index bdc6c55ec4a16..6096792244d94 100644 --- a/Mathlib/Algebra/Squarefree/Basic.lean +++ b/Mathlib/Algebra/Squarefree/Basic.lean @@ -160,9 +160,6 @@ theorem Squarefree.isRadical {x : R} (hx : Squarefree x) : IsRadical x := theorem Squarefree.dvd_pow_iff_dvd {x y : R} {n : ℕ} (hsq : Squarefree x) (h0 : n ≠ 0) : x ∣ y ^ n ↔ x ∣ y := ⟨hsq.isRadical n y, (·.pow h0)⟩ -@[deprecated (since := "2024-02-12")] -alias UniqueFactorizationMonoid.dvd_pow_iff_dvd_of_squarefree := Squarefree.dvd_pow_iff_dvd - end variable [CancelCommMonoidWithZero R] {x y p d : R} diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index 0c8626614a15c..b4eac11407814 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -65,9 +65,6 @@ theorem IsOpenImmersion.isOpen_range {X Y : Scheme.{u}} (f : X ⟶ Y) [H : IsOpe IsOpen (Set.range f.base) := H.base_open.isOpen_range -@[deprecated (since := "2024-03-17")] -alias IsOpenImmersion.open_range := IsOpenImmersion.isOpen_range - namespace Scheme.Hom variable {X Y : Scheme.{u}} (f : Scheme.Hom X Y) [H : IsOpenImmersion f] diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index ba5a97a21ee36..91892e6714502 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -491,7 +491,6 @@ lemma toSpec_fromSpec {f : A} {m : ℕ} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) (x : rw [← FromSpec.num_mem_carrier_iff f_deg hm x] exact ToSpec.mk_mem_carrier _ z -@[deprecated (since := "2024-03-02")] alias toSpecFromSpec := toSpec_fromSpec end toSpecFromSpec diff --git a/Mathlib/Analysis/CStarAlgebra/Matrix.lean b/Mathlib/Analysis/CStarAlgebra/Matrix.lean index 5db2ac0e4cf48..16c1080e57b78 100644 --- a/Mathlib/Analysis/CStarAlgebra/Matrix.lean +++ b/Mathlib/Analysis/CStarAlgebra/Matrix.lean @@ -171,53 +171,35 @@ scoped[Matrix.L2OpNorm] attribute [instance] Matrix.instL2OpNormedAddCommGroup lemma l2_opNorm_def (A : Matrix m n 𝕜) : ‖A‖ = ‖(toEuclideanLin (𝕜 := 𝕜) (m := m) (n := n)).trans toContinuousLinearMap A‖ := rfl -@[deprecated (since := "2024-02-02")] alias l2_op_norm_def := l2_opNorm_def - lemma l2_opNNNorm_def (A : Matrix m n 𝕜) : ‖A‖₊ = ‖(toEuclideanLin (𝕜 := 𝕜) (m := m) (n := n)).trans toContinuousLinearMap A‖₊ := rfl -@[deprecated (since := "2024-02-02")] alias l2_op_nnnorm_def := l2_opNNNorm_def - lemma l2_opNorm_conjTranspose [DecidableEq m] (A : Matrix m n 𝕜) : ‖Aᴴ‖ = ‖A‖ := by rw [l2_opNorm_def, toEuclideanLin_eq_toLin_orthonormal, LinearEquiv.trans_apply, toLin_conjTranspose, adjoint_toContinuousLinearMap] exact ContinuousLinearMap.adjoint.norm_map _ -@[deprecated (since := "2024-02-02")] alias l2_op_norm_conjTranspose := l2_opNorm_conjTranspose - lemma l2_opNNNorm_conjTranspose [DecidableEq m] (A : Matrix m n 𝕜) : ‖Aᴴ‖₊ = ‖A‖₊ := Subtype.ext <| l2_opNorm_conjTranspose _ -@[deprecated (since := "2024-02-02")] alias l2_op_nnnorm_conjTranspose := l2_opNNNorm_conjTranspose - lemma l2_opNorm_conjTranspose_mul_self (A : Matrix m n 𝕜) : ‖Aᴴ * A‖ = ‖A‖ * ‖A‖ := by classical rw [l2_opNorm_def, toEuclideanLin_eq_toLin_orthonormal, LinearEquiv.trans_apply, Matrix.toLin_mul (v₂ := (EuclideanSpace.basisFun m 𝕜).toBasis), toLin_conjTranspose] exact ContinuousLinearMap.norm_adjoint_comp_self _ -@[deprecated (since := "2024-02-02")] -alias l2_op_norm_conjTranspose_mul_self := l2_opNorm_conjTranspose_mul_self - lemma l2_opNNNorm_conjTranspose_mul_self (A : Matrix m n 𝕜) : ‖Aᴴ * A‖₊ = ‖A‖₊ * ‖A‖₊ := Subtype.ext <| l2_opNorm_conjTranspose_mul_self _ -@[deprecated (since := "2024-02-02")] -alias l2_op_nnnorm_conjTranspose_mul_self := l2_opNNNorm_conjTranspose_mul_self - -- note: with only a type ascription in the left-hand side, Lean picks the wrong norm. lemma l2_opNorm_mulVec (A : Matrix m n 𝕜) (x : EuclideanSpace 𝕜 n) : ‖(EuclideanSpace.equiv m 𝕜).symm <| A *ᵥ x‖ ≤ ‖A‖ * ‖x‖ := toEuclideanLin (n := n) (m := m) (𝕜 := 𝕜) |>.trans toContinuousLinearMap A |>.le_opNorm x -@[deprecated (since := "2024-02-02")] alias l2_op_norm_mulVec := l2_opNorm_mulVec - lemma l2_opNNNorm_mulVec (A : Matrix m n 𝕜) (x : EuclideanSpace 𝕜 n) : ‖(EuclideanSpace.equiv m 𝕜).symm <| A *ᵥ x‖₊ ≤ ‖A‖₊ * ‖x‖₊ := A.l2_opNorm_mulVec x -@[deprecated (since := "2024-02-02")] alias l2_op_nnnorm_mulVec := l2_opNNNorm_mulVec - lemma l2_opNorm_mul (A : Matrix m n 𝕜) (B : Matrix n l 𝕜) : ‖A * B‖ ≤ ‖A‖ * ‖B‖ := by simp only [l2_opNorm_def] @@ -227,13 +209,9 @@ lemma l2_opNorm_mul (A : Matrix m n 𝕜) (B : Matrix n l 𝕜) : ext1 x exact congr($(Matrix.toLin'_mul A B) x) -@[deprecated (since := "2024-02-02")] alias l2_op_norm_mul := l2_opNorm_mul - lemma l2_opNNNorm_mul (A : Matrix m n 𝕜) (B : Matrix n l 𝕜) : ‖A * B‖₊ ≤ ‖A‖₊ * ‖B‖₊ := l2_opNorm_mul A B -@[deprecated (since := "2024-02-02")] alias l2_op_nnnorm_mul := l2_opNNNorm_mul - /-- The normed algebra structure on `Matrix n n 𝕜` arising from the operator norm given by the identification with (continuous) linear endmorphisms of `EuclideanSpace 𝕜 n`. -/ def instL2OpNormedSpace : NormedSpace 𝕜 (Matrix m n 𝕜) where diff --git a/Mathlib/Analysis/CStarAlgebra/Unitization.lean b/Mathlib/Analysis/CStarAlgebra/Unitization.lean index 8771d6ba22cae..caa411ee70756 100644 --- a/Mathlib/Analysis/CStarAlgebra/Unitization.lean +++ b/Mathlib/Analysis/CStarAlgebra/Unitization.lean @@ -38,13 +38,9 @@ lemma opNorm_mul_flip_apply (a : E) : ‖(mul 𝕜 E).flip a‖ = ‖a‖ := by _ ≤ ‖(mul 𝕜 E).flip a‖ * ‖b‖ := by simpa only [flip_apply, mul_apply', norm_star] using le_opNorm ((mul 𝕜 E).flip a) (star b) -@[deprecated (since := "2024-02-02")] alias op_norm_mul_flip_apply := opNorm_mul_flip_apply - lemma opNNNorm_mul_flip_apply (a : E) : ‖(mul 𝕜 E).flip a‖₊ = ‖a‖₊ := Subtype.ext (opNorm_mul_flip_apply 𝕜 a) -@[deprecated (since := "2024-02-02")] alias op_nnnorm_mul_flip_apply := opNNNorm_mul_flip_apply - variable (E) lemma isometry_mul_flip : Isometry (mul 𝕜 E).flip := diff --git a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean index 400df9cb52915..168c48c5ac9ff 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean @@ -630,10 +630,6 @@ theorem HasFTaylorSeriesUpToOn.eq_iteratedFDerivWithin_of_uniqueDiffOn rw [iteratedFDerivWithin_succ_eq_comp_left, Function.comp_apply, this.fderivWithin (hs x hx)] exact (ContinuousMultilinearMap.uncurry_curryLeft _).symm -@[deprecated (since := "2024-03-28")] -alias HasFTaylorSeriesUpToOn.eq_ftaylor_series_of_uniqueDiffOn := - HasFTaylorSeriesUpToOn.eq_iteratedFDerivWithin_of_uniqueDiffOn - /-- The iterated derivative commutes with shifting the function by a constant on the left. -/ lemma iteratedFDerivWithin_comp_add_left' (n : ℕ) (a : E) : iteratedFDerivWithin 𝕜 n (fun z ↦ f (a + z)) s = diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean index 48cbff9518b9a..be38e1e55e5cc 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean @@ -49,5 +49,3 @@ variable {f} theorem isOpenMap_of_hasStrictDerivAt {f' : 𝕜 → 𝕜} (hf : ∀ x, HasStrictDerivAt f (f' x) x) (h0 : ∀ x, f' x ≠ 0) : IsOpenMap f := isOpenMap_iff_nhds_le.2 fun x => ((hf x).map_nhds_eq (h0 x)).ge -@[deprecated (since := "2024-03-23")] -alias open_map_of_strict_deriv := isOpenMap_of_hasStrictDerivAt diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean index 06f438efc45a0..7f5ef4e24ef29 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean @@ -194,5 +194,3 @@ end HasStrictFDerivAt theorem isOpenMap_of_hasStrictFDerivAt_equiv [CompleteSpace E] {f : E → F} {f' : E → E ≃L[𝕜] F} (hf : ∀ x, HasStrictFDerivAt f (f' x : E →L[𝕜] F) x) : IsOpenMap f := isOpenMap_iff_nhds_le.2 fun x => (hf x).map_nhds_eq_of_equiv.ge -@[deprecated (since := "2024-03-23")] -alias open_map_of_strict_fderiv_equiv := isOpenMap_of_hasStrictFDerivAt_equiv diff --git a/Mathlib/Analysis/Calculus/MeanValue.lean b/Mathlib/Analysis/Calculus/MeanValue.lean index 8194a4e3aa8f5..ee6d474ded239 100644 --- a/Mathlib/Analysis/Calculus/MeanValue.lean +++ b/Mathlib/Analysis/Calculus/MeanValue.lean @@ -1008,9 +1008,6 @@ lemma strictMonoOn_of_hasDerivWithinAt_pos {D : Set ℝ} (hD : Convex ℝ D) {f strictMonoOn_of_deriv_pos hD hf fun x hx ↦ by rw [deriv_eqOn isOpen_interior hf' hx]; exact hf'₀ _ hx -@[deprecated (since := "2024-03-02")] -alias StrictMonoOn_of_hasDerivWithinAt_pos := strictMonoOn_of_hasDerivWithinAt_pos - /-- Let `f : ℝ → ℝ` be a differentiable function. If `f'` is strictly positive, then `f` is a strictly monotone function. -/ lemma strictMono_of_hasDerivAt_pos {f f' : ℝ → ℝ} (hf : ∀ x, HasDerivAt f (f' x) x) @@ -1080,18 +1077,12 @@ lemma strictAntiOn_of_hasDerivWithinAt_neg {D : Set ℝ} (hD : Convex ℝ D) {f strictAntiOn_of_deriv_neg hD hf fun x hx ↦ by rw [deriv_eqOn isOpen_interior hf' hx]; exact hf'₀ _ hx -@[deprecated (since := "2024-03-02")] -alias StrictAntiOn_of_hasDerivWithinAt_pos := strictAntiOn_of_hasDerivWithinAt_neg - /-- Let `f : ℝ → ℝ` be a differentiable function. If `f'` is strictly positive, then `f` is a strictly monotone function. -/ lemma strictAnti_of_hasDerivAt_neg {f f' : ℝ → ℝ} (hf : ∀ x, HasDerivAt f (f' x) x) (hf' : ∀ x, f' x < 0) : StrictAnti f := strictAnti_of_deriv_neg fun x ↦ by rw [(hf _).deriv]; exact hf' _ -@[deprecated (since := "2024-03-02")] -alias strictAnti_of_hasDerivAt_pos := strictAnti_of_hasDerivAt_neg - /-- Let `f` be a function continuous on a convex (or, equivalently, connected) subset `D` of the real line. If `f` is differentiable on the interior of `D` and `f'` is nonpositive, then `f` is an antitone function on `D`. -/ diff --git a/Mathlib/Analysis/Complex/Polynomial/Basic.lean b/Mathlib/Analysis/Complex/Polynomial/Basic.lean index 82e38a8adc578..706b89c25be7f 100644 --- a/Mathlib/Analysis/Complex/Polynomial/Basic.lean +++ b/Mathlib/Analysis/Complex/Polynomial/Basic.lean @@ -216,6 +216,3 @@ lemma Irreducible.degree_le_two {p : ℝ[X]} (hp : Irreducible p) : degree p ≤ /-- An irreducible real polynomial has natural degree at most two. -/ lemma Irreducible.natDegree_le_two {p : ℝ[X]} (hp : Irreducible p) : natDegree p ≤ 2 := natDegree_le_iff_degree_le.2 hp.degree_le_two - -@[deprecated (since := "2024-02-18")] -alias Irreducible.nat_degree_le_two := Irreducible.natDegree_le_two diff --git a/Mathlib/Analysis/Fourier/FourierTransform.lean b/Mathlib/Analysis/Fourier/FourierTransform.lean index a5d1ddb0067c9..c75680f9ddcd1 100644 --- a/Mathlib/Analysis/Fourier/FourierTransform.lean +++ b/Mathlib/Analysis/Fourier/FourierTransform.lean @@ -136,9 +136,6 @@ theorem fourierIntegral_convergent_iff (he : Continuous e) e.map_zero_eq_one, one_smul] at this -- the `(e _)` speeds up elaboration considerably exact this -@[deprecated (since := "2024-03-29")] -alias fourier_integral_convergent_iff := VectorFourier.fourierIntegral_convergent_iff - 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 + g) = fourierIntegral e μ L f + fourierIntegral e μ L g := by @@ -431,16 +428,11 @@ theorem fourierIntegral_real_eq (f : ℝ → E) (w : ℝ) : fourierIntegral f w = ∫ v : ℝ, 𝐞 (-(v * w)) • f v := rfl -@[deprecated (since := "2024-02-21")] alias fourierIntegral_def := fourierIntegral_real_eq - 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, Circle.smul_def, Real.fourierChar_apply, mul_neg, neg_mul, mul_assoc] -@[deprecated (since := "2024-02-21")] -alias fourierIntegral_eq_integral_exp_smul := fourierIntegral_real_eq_integral_exp_smul - 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/InnerProductSpace/LaxMilgram.lean b/Mathlib/Analysis/InnerProductSpace/LaxMilgram.lean index 9aa623e738fae..4e702fd38f691 100644 --- a/Mathlib/Analysis/InnerProductSpace/LaxMilgram.lean +++ b/Mathlib/Analysis/InnerProductSpace/LaxMilgram.lean @@ -76,7 +76,6 @@ theorem isClosed_range (coercive : IsCoercive B) : IsClosed (range B♯ : Set V) rcases coercive.antilipschitz with ⟨_, _, antilipschitz⟩ exact antilipschitz.isClosed_range B♯.uniformContinuous -@[deprecated (since := "2024-03-19")] alias closed_range := isClosed_range theorem range_eq_top (coercive : IsCoercive B) : range B♯ = ⊤ := by haveI := coercive.isClosed_range.completeSpace_coe diff --git a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean index a9afb01f2a21b..ce34bcdfe2375 100644 --- a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean +++ b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean @@ -79,10 +79,6 @@ lemma FiniteDimensional.of_fact_finrank_eq_two {K V : Type*} [DivisionRing K] attribute [local instance] FiniteDimensional.of_fact_finrank_eq_two -@[deprecated (since := "2024-02-02")] -alias FiniteDimensional.finiteDimensional_of_fact_finrank_eq_two := - FiniteDimensional.of_fact_finrank_eq_two - variable {E : Type*} [NormedAddCommGroup E] [InnerProductSpace ℝ E] [Fact (finrank ℝ E = 2)] (o : Orientation ℝ E (Fin 2)) diff --git a/Mathlib/Analysis/LocallyConvex/Basic.lean b/Mathlib/Analysis/LocallyConvex/Basic.lean index ffe0e4d997485..f6d8feaab62d1 100644 --- a/Mathlib/Analysis/LocallyConvex/Basic.lean +++ b/Mathlib/Analysis/LocallyConvex/Basic.lean @@ -215,8 +215,6 @@ theorem Balanced.absorbs_self (hA : Balanced 𝕜 A) : Absorbs 𝕜 A A := theorem Balanced.smul_mem_iff (hs : Balanced 𝕜 s) (h : ‖a‖ = ‖b‖) : a • x ∈ s ↔ b • x ∈ s := ⟨(hs.smul_mem_mono · h.ge), (hs.smul_mem_mono · h.le)⟩ -@[deprecated (since := "2024-02-02")] alias Balanced.mem_smul_iff := Balanced.smul_mem_iff - variable [TopologicalSpace E] [ContinuousSMul 𝕜 E] /-- Every neighbourhood of the origin is absorbent. -/ @@ -235,10 +233,6 @@ theorem Balanced.zero_insert_interior (hA : Balanced 𝕜 A) : apply insert_subset_insert exact ((isOpenMap_smul₀ h).mapsTo_interior <| hA.smul_mem ha).image_subset -@[deprecated Balanced.zero_insert_interior (since := "2024-02-03")] -theorem balanced_zero_union_interior (hA : Balanced 𝕜 A) : Balanced 𝕜 ((0 : Set E) ∪ interior A) := - hA.zero_insert_interior - /-- The interior of a balanced set is balanced if it contains the origin. -/ protected theorem Balanced.interior (hA : Balanced 𝕜 A) (h : (0 : E) ∈ interior A) : Balanced 𝕜 (interior A) := by @@ -255,9 +249,6 @@ section NontriviallyNormedField variable [NontriviallyNormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} -@[deprecated Absorbent.zero_mem (since := "2024-02-02")] -theorem Absorbent.zero_mem' (hs : Absorbent 𝕜 s) : (0 : E) ∈ s := hs.zero_mem - variable [Module ℝ E] [SMulCommClass ℝ 𝕜 E] protected theorem Balanced.convexHull (hs : Balanced 𝕜 s) : Balanced 𝕜 (convexHull ℝ s) := by @@ -269,8 +260,6 @@ protected theorem Balanced.convexHull (hs : Balanced 𝕜 s) : Balanced 𝕜 (co simp only [smul_add, ← smul_comm] exact convex_convexHull ℝ s (hx a ha) (hy a ha) hu hv huv -@[deprecated (since := "2024-02-02")] alias balanced_convexHull_of_balanced := Balanced.convexHull - end NontriviallyNormedField section Real diff --git a/Mathlib/Analysis/Matrix.lean b/Mathlib/Analysis/Matrix.lean index 7a1f22fb8be30..d82d5f067fe3c 100644 --- a/Mathlib/Analysis/Matrix.lean +++ b/Mathlib/Analysis/Matrix.lean @@ -248,39 +248,27 @@ theorem linfty_opNorm_def (A : Matrix m n α) : change ‖fun i => (WithLp.equiv 1 _).symm (A i)‖ = _ simp [Pi.norm_def, PiLp.nnnorm_eq_of_L1] -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_def := linfty_opNorm_def - theorem linfty_opNNNorm_def (A : Matrix m n α) : ‖A‖₊ = (Finset.univ : Finset m).sup fun i : m => ∑ j : n, ‖A i j‖₊ := Subtype.ext <| linfty_opNorm_def A -@[deprecated (since := "2024-02-02")] alias linfty_op_nnnorm_def := linfty_opNNNorm_def - @[simp] theorem linfty_opNNNorm_col (v : m → α) : ‖col ι v‖₊ = ‖v‖₊ := by rw [linfty_opNNNorm_def, Pi.nnnorm_def] simp -@[deprecated (since := "2024-02-02")] alias linfty_op_nnnorm_col := linfty_opNNNorm_col - @[simp] theorem linfty_opNorm_col (v : m → α) : ‖col ι v‖ = ‖v‖ := congr_arg ((↑) : ℝ≥0 → ℝ) <| linfty_opNNNorm_col v -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_col := linfty_opNorm_col - @[simp] theorem linfty_opNNNorm_row (v : n → α) : ‖row ι v‖₊ = ∑ i, ‖v i‖₊ := by simp [linfty_opNNNorm_def] -@[deprecated (since := "2024-02-02")] alias linfty_op_nnnorm_row := linfty_opNNNorm_row - @[simp] theorem linfty_opNorm_row (v : n → α) : ‖row ι v‖ = ∑ i, ‖v i‖ := (congr_arg ((↑) : ℝ≥0 → ℝ) <| linfty_opNNNorm_row v).trans <| by simp [NNReal.coe_sum] -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_row := linfty_opNorm_row - @[simp] theorem linfty_opNNNorm_diagonal [DecidableEq m] (v : m → α) : ‖diagonal v‖₊ = ‖v‖₊ := by rw [linfty_opNNNorm_def, Pi.nnnorm_def] @@ -289,14 +277,10 @@ theorem linfty_opNNNorm_diagonal [DecidableEq m] (v : m → α) : ‖diagonal v · rw [diagonal_apply_ne' _ hij, nnnorm_zero] · rw [diagonal_apply_eq] -@[deprecated (since := "2024-02-02")] alias linfty_op_nnnorm_diagonal := linfty_opNNNorm_diagonal - @[simp] theorem linfty_opNorm_diagonal [DecidableEq m] (v : m → α) : ‖diagonal v‖ = ‖v‖ := congr_arg ((↑) : ℝ≥0 → ℝ) <| linfty_opNNNorm_diagonal v -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_diagonal := linfty_opNorm_diagonal - end SeminormedAddCommGroup section NonUnitalSeminormedRing @@ -320,24 +304,16 @@ theorem linfty_opNNNorm_mul (A : Matrix l m α) (B : Matrix m n α) : ‖A * B simp_rw [← Finset.sum_mul, ← NNReal.finset_sup_mul] rfl -@[deprecated (since := "2024-02-02")] alias linfty_op_nnnorm_mul := linfty_opNNNorm_mul - theorem linfty_opNorm_mul (A : Matrix l m α) (B : Matrix m n α) : ‖A * B‖ ≤ ‖A‖ * ‖B‖ := linfty_opNNNorm_mul _ _ -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_mul := linfty_opNorm_mul - theorem linfty_opNNNorm_mulVec (A : Matrix l m α) (v : m → α) : ‖A *ᵥ v‖₊ ≤ ‖A‖₊ * ‖v‖₊ := by rw [← linfty_opNNNorm_col (ι := Fin 1) (A *ᵥ v), ← linfty_opNNNorm_col v (ι := Fin 1)] exact linfty_opNNNorm_mul A (col (Fin 1) v) -@[deprecated (since := "2024-02-02")] alias linfty_op_nnnorm_mulVec := linfty_opNNNorm_mulVec - theorem linfty_opNorm_mulVec (A : Matrix l m α) (v : m → α) : ‖A *ᵥ v‖ ≤ ‖A‖ * ‖v‖ := linfty_opNNNorm_mulVec _ _ -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_mulVec := linfty_opNorm_mulVec - end NonUnitalSeminormedRing /-- Seminormed non-unital ring instance (using sup norm of L1 norm) for matrices over a semi normed @@ -435,15 +411,10 @@ lemma linfty_opNNNorm_eq_opNNNorm (A : Matrix m n α) : nnnorm_one, mul_one] at hN exact hN -@[deprecated (since := "2024-02-02")] -alias linfty_op_nnnorm_eq_op_nnnorm := linfty_opNNNorm_eq_opNNNorm - lemma linfty_opNorm_eq_opNorm (A : Matrix m n α) : ‖A‖ = ‖ContinuousLinearMap.mk (Matrix.mulVecLin A)‖ := congr_arg NNReal.toReal (linfty_opNNNorm_eq_opNNNorm A) -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_eq_op_norm := linfty_opNorm_eq_opNorm - variable [DecidableEq n] @[simp] lemma linfty_opNNNorm_toMatrix (f : (n → α) →L[α] (m → α)) : @@ -451,14 +422,10 @@ variable [DecidableEq n] rw [linfty_opNNNorm_eq_opNNNorm] simp only [← toLin'_apply', toLin'_toMatrix'] -@[deprecated (since := "2024-02-02")] alias linfty_op_nnnorm_toMatrix := linfty_opNNNorm_toMatrix - @[simp] lemma linfty_opNorm_toMatrix (f : (n → α) →L[α] (m → α)) : ‖LinearMap.toMatrix' (↑f : (n → α) →ₗ[α] (m → α))‖ = ‖f‖ := congr_arg NNReal.toReal (linfty_opNNNorm_toMatrix f) -@[deprecated (since := "2024-02-02")] alias linfty_op_norm_toMatrix := linfty_opNorm_toMatrix - end end LinftyOp diff --git a/Mathlib/Analysis/Normed/Lp/LpEquiv.lean b/Mathlib/Analysis/Normed/Lp/LpEquiv.lean index d77bc7fea6e3e..bd8818ac7ee46 100644 --- a/Mathlib/Analysis/Normed/Lp/LpEquiv.lean +++ b/Mathlib/Analysis/Normed/Lp/LpEquiv.lean @@ -133,7 +133,6 @@ noncomputable def AddEquiv.lpBCF : lp (fun _ : α ↦ E) ∞ ≃+ (α →ᵇ E) right_inv _f := rfl map_add' _f _g := rfl -@[deprecated (since := "2024-03-16")] alias AddEquiv.lpBcf := AddEquiv.lpBCF theorem coe_addEquiv_lpBCF (f : lp (fun _ : α ↦ E) ∞) : (AddEquiv.lpBCF f : α → E) = f := rfl @@ -149,7 +148,6 @@ noncomputable def lpBCFₗᵢ : lp (fun _ : α ↦ E) ∞ ≃ₗᵢ[𝕜] α → map_smul' := fun _ _ ↦ rfl norm_map' := fun f ↦ by simp only [norm_eq_iSup_norm, lp.norm_eq_ciSup]; rfl } -@[deprecated (since := "2024-03-16")] alias lpBcfₗᵢ := lpBCFₗᵢ variable {𝕜 E} @@ -168,7 +166,6 @@ noncomputable def RingEquiv.lpBCF : lp (fun _ : α ↦ R) ∞ ≃+* (α →ᵇ R { @AddEquiv.lpBCF _ R _ _ _ with map_mul' := fun _f _g => rfl } -@[deprecated (since := "2024-03-16")] alias RingEquiv.lpBcf := RingEquiv.lpBCF variable {R} @@ -187,7 +184,6 @@ variable (α) noncomputable def AlgEquiv.lpBCF : lp (fun _ : α ↦ A) ∞ ≃ₐ[𝕜] α →ᵇ A := { RingEquiv.lpBCF A with commutes' := fun _k ↦ rfl } -@[deprecated (since := "2024-03-16")] alias AlgEquiv.lpBcf := AlgEquiv.lpBCF variable {α A 𝕜} diff --git a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean index 789c690f2042f..3620ce5749a6c 100644 --- a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean +++ b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean @@ -287,15 +287,11 @@ theorem Basis.opNNNorm_le {ι : Type*} [Fintype ι] (v : Basis ι 𝕜 E) {u : E _ ≤ Fintype.card ι • (‖φ‖₊ * ‖e‖₊) := nsmul_le_nsmul_right (φ.le_opNNNorm e) _ _ = Fintype.card ι • ‖φ‖₊ * M * ‖e‖₊ := by simp only [smul_mul_assoc, mul_right_comm] -@[deprecated (since := "2024-02-02")] alias Basis.op_nnnorm_le := Basis.opNNNorm_le - theorem Basis.opNorm_le {ι : Type*} [Fintype ι] (v : Basis ι 𝕜 E) {u : E →L[𝕜] F} {M : ℝ} (hM : 0 ≤ M) (hu : ∀ i, ‖u (v i)‖ ≤ M) : ‖u‖ ≤ Fintype.card ι • ‖v.equivFunL.toContinuousLinearMap‖ * M := by simpa using NNReal.coe_le_coe.mpr (v.opNNNorm_le ⟨M, hM⟩ hu) -@[deprecated (since := "2024-02-02")] alias Basis.op_norm_le := Basis.opNorm_le - /-- A weaker version of `Basis.opNNNorm_le` that abstracts away the value of `C`. -/ theorem Basis.exists_opNNNorm_le {ι : Type*} [Finite ι] (v : Basis ι 𝕜 E) : ∃ C > (0 : ℝ≥0), ∀ {u : E →L[𝕜] F} (M : ℝ≥0), (∀ i, ‖u (v i)‖₊ ≤ M) → ‖u‖₊ ≤ C * M := by @@ -305,8 +301,6 @@ theorem Basis.exists_opNNNorm_le {ι : Type*} [Finite ι] (v : Basis ι 𝕜 E) zero_lt_one.trans_le (le_max_right _ _), fun {u} M hu => (v.opNNNorm_le M hu).trans <| mul_le_mul_of_nonneg_right (le_max_left _ _) (zero_le M)⟩ -@[deprecated (since := "2024-02-02")] alias Basis.exists_op_nnnorm_le := Basis.exists_opNNNorm_le - /-- A weaker version of `Basis.opNorm_le` that abstracts away the value of `C`. -/ theorem Basis.exists_opNorm_le {ι : Type*} [Finite ι] (v : Basis ι 𝕜 E) : ∃ C > (0 : ℝ), ∀ {u : E →L[𝕜] F} {M : ℝ}, 0 ≤ M → (∀ i, ‖u (v i)‖ ≤ M) → ‖u‖ ≤ C * M := by @@ -316,8 +310,6 @@ theorem Basis.exists_opNorm_le {ι : Type*} [Finite ι] (v : Basis ι 𝕜 E) : intro u M hM H simpa using h ⟨M, hM⟩ H -@[deprecated (since := "2024-02-02")] alias Basis.exists_op_norm_le := Basis.exists_opNorm_le - instance [FiniteDimensional 𝕜 E] [SecondCountableTopology F] : SecondCountableTopology (E →L[𝕜] F) := by set d := Module.finrank 𝕜 E @@ -455,27 +447,18 @@ theorem FiniteDimensional.of_isCompact_closedBall₀ {r : ℝ} (rpos : 0 < r) exact φmono (Nat.lt_succ_self N) _ < ‖c‖ := hN (N + 1) (Nat.le_succ N) -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_isCompact_closedBall₀ := FiniteDimensional.of_isCompact_closedBall₀ - /-- **Riesz's theorem**: if a closed ball of positive radius is compact in a vector space, then the space is finite-dimensional. -/ theorem FiniteDimensional.of_isCompact_closedBall {r : ℝ} (rpos : 0 < r) {c : E} (h : IsCompact (Metric.closedBall c r)) : FiniteDimensional 𝕜 E := .of_isCompact_closedBall₀ 𝕜 rpos <| by simpa using h.vadd (-c) -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_isCompact_closedBall := FiniteDimensional.of_isCompact_closedBall - /-- **Riesz's theorem**: a locally compact normed vector space is finite-dimensional. -/ theorem FiniteDimensional.of_locallyCompactSpace [LocallyCompactSpace E] : FiniteDimensional 𝕜 E := let ⟨_r, rpos, hr⟩ := exists_isCompact_closedBall (0 : E) .of_isCompact_closedBall₀ 𝕜 rpos hr -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_locallyCompactSpace := FiniteDimensional.of_locallyCompactSpace - /-- If a function has compact support, then either the function is trivial or the space is finite-dimensional. -/ theorem HasCompactSupport.eq_zero_or_finiteDimensional {X : Type*} [TopologicalSpace X] [Zero X] diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 4a71d1dff9424..e45477c7b3695 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -374,29 +374,20 @@ theorem isLeast_opNorm (f : ContinuousMultilinearMap 𝕜 E G) : exact isClosed_Ici.inter (isClosed_iInter fun m ↦ isClosed_le continuous_const (continuous_id.mul continuous_const)) -@[deprecated (since := "2024-02-02")] alias isLeast_op_norm := isLeast_opNorm - theorem opNorm_nonneg (f : ContinuousMultilinearMap 𝕜 E G) : 0 ≤ ‖f‖ := Real.sInf_nonneg fun _ ⟨hx, _⟩ => hx -@[deprecated (since := "2024-02-02")] alias op_norm_nonneg := opNorm_nonneg - /-- The fundamental property of the operator norm of a continuous multilinear map: `‖f m‖` is bounded by `‖f‖` times the product of the `‖m i‖`. -/ theorem le_opNorm (f : ContinuousMultilinearMap 𝕜 E G) (m : ∀ i, E i) : ‖f m‖ ≤ ‖f‖ * ∏ i, ‖m i‖ := f.isLeast_opNorm.1.2 m -@[deprecated (since := "2024-02-02")] alias le_op_norm := le_opNorm - theorem le_mul_prod_of_opNorm_le_of_le {f : ContinuousMultilinearMap 𝕜 E G} {m : ∀ i, E i} {C : ℝ} {b : ι → ℝ} (hC : ‖f‖ ≤ C) (hm : ∀ i, ‖m i‖ ≤ b i) : ‖f m‖ ≤ C * ∏ i, b i := (f.le_opNorm m).trans <| by gcongr; exacts [f.opNorm_nonneg.trans hC, hm _] -@[deprecated (since := "2024-02-02")] -alias le_mul_prod_of_le_op_norm_of_le := le_mul_prod_of_opNorm_le_of_le - @[deprecated (since := "2024-11-27")] alias le_mul_prod_of_le_opNorm_of_le := le_mul_prod_of_opNorm_le_of_le @@ -404,67 +395,46 @@ theorem le_opNorm_mul_prod_of_le (f : ContinuousMultilinearMap 𝕜 E G) {m : ∀ i, E i} {b : ι → ℝ} (hm : ∀ i, ‖m i‖ ≤ b i) : ‖f m‖ ≤ ‖f‖ * ∏ i, b i := le_mul_prod_of_opNorm_le_of_le le_rfl hm -@[deprecated (since := "2024-02-02")] alias le_op_norm_mul_prod_of_le := le_opNorm_mul_prod_of_le - theorem le_opNorm_mul_pow_card_of_le (f : ContinuousMultilinearMap 𝕜 E G) {m b} (hm : ‖m‖ ≤ b) : ‖f m‖ ≤ ‖f‖ * b ^ Fintype.card ι := by simpa only [prod_const] using f.le_opNorm_mul_prod_of_le fun i => (norm_le_pi_norm m i).trans hm -@[deprecated (since := "2024-02-02")] -alias le_op_norm_mul_pow_card_of_le := le_opNorm_mul_pow_card_of_le - theorem le_opNorm_mul_pow_of_le {n : ℕ} {Ei : Fin n → Type*} [∀ i, SeminormedAddCommGroup (Ei i)] [∀ i, NormedSpace 𝕜 (Ei i)] (f : ContinuousMultilinearMap 𝕜 Ei G) {m : ∀ i, Ei i} {b : ℝ} (hm : ‖m‖ ≤ b) : ‖f m‖ ≤ ‖f‖ * b ^ n := by simpa only [Fintype.card_fin] using f.le_opNorm_mul_pow_card_of_le hm -@[deprecated (since := "2024-02-02")] alias le_op_norm_mul_pow_of_le := le_opNorm_mul_pow_of_le - theorem le_of_opNorm_le {f : ContinuousMultilinearMap 𝕜 E G} {C : ℝ} (h : ‖f‖ ≤ C) (m : ∀ i, E i) : ‖f m‖ ≤ C * ∏ i, ‖m i‖ := le_mul_prod_of_opNorm_le_of_le h fun _ ↦ le_rfl -@[deprecated (since := "2024-02-02")] alias le_of_op_norm_le := le_of_opNorm_le - theorem ratio_le_opNorm (f : ContinuousMultilinearMap 𝕜 E G) (m : ∀ i, E i) : (‖f m‖ / ∏ i, ‖m i‖) ≤ ‖f‖ := div_le_of_le_mul₀ (by positivity) (opNorm_nonneg _) (f.le_opNorm m) -@[deprecated (since := "2024-02-02")] alias ratio_le_op_norm := ratio_le_opNorm - /-- The image of the unit ball under a continuous multilinear map is bounded. -/ theorem unit_le_opNorm (f : ContinuousMultilinearMap 𝕜 E G) {m : ∀ i, E i} (h : ‖m‖ ≤ 1) : ‖f m‖ ≤ ‖f‖ := (le_opNorm_mul_pow_card_of_le f h).trans <| by simp -@[deprecated (since := "2024-02-02")] alias unit_le_op_norm := unit_le_opNorm - /-- If one controls the norm of every `f x`, then one controls the norm of `f`. -/ theorem opNorm_le_bound {f : ContinuousMultilinearMap 𝕜 E G} {M : ℝ} (hMp : 0 ≤ M) (hM : ∀ m, ‖f m‖ ≤ M * ∏ i, ‖m i‖) : ‖f‖ ≤ M := csInf_le bounds_bddBelow ⟨hMp, hM⟩ -@[deprecated (since := "2024-02-02")] alias op_norm_le_bound := opNorm_le_bound - theorem opNorm_le_iff {f : ContinuousMultilinearMap 𝕜 E G} {C : ℝ} (hC : 0 ≤ C) : ‖f‖ ≤ C ↔ ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖ := ⟨fun h _ ↦ le_of_opNorm_le h _, opNorm_le_bound hC⟩ -@[deprecated (since := "2024-02-02")] alias op_norm_le_iff := opNorm_le_iff - /-- The operator norm satisfies the triangle inequality. -/ theorem opNorm_add_le (f g : ContinuousMultilinearMap 𝕜 E G) : ‖f + g‖ ≤ ‖f‖ + ‖g‖ := opNorm_le_bound (add_nonneg (opNorm_nonneg f) (opNorm_nonneg g)) fun x => by rw [add_mul] exact norm_add_le_of_le (le_opNorm _ _) (le_opNorm _ _) -@[deprecated (since := "2024-02-02")] alias op_norm_add_le := opNorm_add_le - theorem opNorm_zero : ‖(0 : ContinuousMultilinearMap 𝕜 E G)‖ = 0 := (opNorm_nonneg _).antisymm' <| opNorm_le_bound le_rfl fun m => by simp -@[deprecated (since := "2024-02-02")] alias op_norm_zero := opNorm_zero - section variable {𝕜' : Type*} [NormedField 𝕜'] [NormedSpace 𝕜' G] [SMulCommClass 𝕜 𝕜' G] @@ -474,8 +444,6 @@ theorem opNorm_smul_le (c : 𝕜') (f : ContinuousMultilinearMap 𝕜 E G) : ‖ rw [smul_apply, norm_smul, mul_assoc] exact mul_le_mul_of_nonneg_left (le_opNorm _ _) (norm_nonneg _) -@[deprecated (since := "2024-02-02")] alias op_norm_smul_le := opNorm_smul_le - variable (𝕜 E G) in /-- Operator seminorm on the space of continuous multilinear maps, as `Seminorm`. @@ -544,8 +512,6 @@ instance normedSpace' : NormedSpace 𝕜' (ContinuousMultilinearMap 𝕜 (fun _ @[deprecated norm_neg (since := "2024-11-24")] theorem opNorm_neg (f : ContinuousMultilinearMap 𝕜 E G) : ‖-f‖ = ‖f‖ := norm_neg f -@[deprecated (since := "2024-02-02")] alias op_norm_neg := norm_neg - /-- The fundamental property of the operator norm of a continuous multilinear map: `‖f m‖` is bounded by `‖f‖` times the product of the `‖m i‖`, `nnnorm` version. -/ theorem le_opNNNorm (f : ContinuousMultilinearMap 𝕜 E G) (m : ∀ i, E i) : @@ -554,39 +520,27 @@ theorem le_opNNNorm (f : ContinuousMultilinearMap 𝕜 E G) (m : ∀ i, E i) : push_cast exact f.le_opNorm m -@[deprecated (since := "2024-02-02")] alias le_op_nnnorm := le_opNNNorm - theorem le_of_opNNNorm_le (f : ContinuousMultilinearMap 𝕜 E G) {C : ℝ≥0} (h : ‖f‖₊ ≤ C) (m : ∀ i, E i) : ‖f m‖₊ ≤ C * ∏ i, ‖m i‖₊ := (f.le_opNNNorm m).trans <| mul_le_mul' h le_rfl -@[deprecated (since := "2024-02-02")] alias le_of_op_nnnorm_le := le_of_opNNNorm_le - theorem opNNNorm_le_iff {f : ContinuousMultilinearMap 𝕜 E G} {C : ℝ≥0} : ‖f‖₊ ≤ C ↔ ∀ m, ‖f m‖₊ ≤ C * ∏ i, ‖m i‖₊ := by simp only [← NNReal.coe_le_coe]; simp [opNorm_le_iff C.coe_nonneg, NNReal.coe_prod] -@[deprecated (since := "2024-02-02")] alias op_nnnorm_le_iff := opNNNorm_le_iff - theorem isLeast_opNNNorm (f : ContinuousMultilinearMap 𝕜 E G) : IsLeast {C : ℝ≥0 | ∀ m, ‖f m‖₊ ≤ C * ∏ i, ‖m i‖₊} ‖f‖₊ := by simpa only [← opNNNorm_le_iff] using isLeast_Ici -@[deprecated (since := "2024-02-02")] alias isLeast_op_nnnorm := isLeast_opNNNorm - theorem opNNNorm_prod (f : ContinuousMultilinearMap 𝕜 E G) (g : ContinuousMultilinearMap 𝕜 E G') : ‖f.prod g‖₊ = max ‖f‖₊ ‖g‖₊ := eq_of_forall_ge_iff fun _ ↦ by simp only [opNNNorm_le_iff, prod_apply, Prod.nnnorm_def', max_le_iff, forall_and] -@[deprecated (since := "2024-02-02")] alias op_nnnorm_prod := opNNNorm_prod - theorem opNorm_prod (f : ContinuousMultilinearMap 𝕜 E G) (g : ContinuousMultilinearMap 𝕜 E G') : ‖f.prod g‖ = max ‖f‖ ‖g‖ := congr_arg NNReal.toReal (opNNNorm_prod f g) -@[deprecated (since := "2024-02-02")] alias op_norm_prod := opNorm_prod - theorem opNNNorm_pi [∀ i', SeminormedAddCommGroup (E' i')] [∀ i', NormedSpace 𝕜 (E' i')] (f : ∀ i', ContinuousMultilinearMap 𝕜 E (E' i')) : ‖pi f‖₊ = ‖f‖₊ := @@ -598,8 +552,6 @@ theorem opNorm_pi {ι' : Type v'} [Fintype ι'] {E' : ι' → Type wE'} ‖pi f‖ = ‖f‖ := congr_arg NNReal.toReal (opNNNorm_pi f) -@[deprecated (since := "2024-02-02")] alias op_norm_pi := opNorm_pi - section @[simp] @@ -1332,8 +1284,6 @@ variable {𝕜 : Type u} {ι : Type v} {E : ι → Type wE} {G : Type wG} {G' : theorem opNorm_zero_iff {f : ContinuousMultilinearMap 𝕜 E G} : ‖f‖ = 0 ↔ f = 0 := by simp [← (opNorm_nonneg f).le_iff_eq, opNorm_le_iff le_rfl, ContinuousMultilinearMap.ext_iff] -@[deprecated (since := "2024-02-02")] alias op_norm_zero_iff := opNorm_zero_iff - /-- Continuous multilinear maps themselves form a normed group with respect to the operator norm. -/ instance normedAddCommGroup : NormedAddCommGroup (ContinuousMultilinearMap 𝕜 E G) := diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean index 34ac1d9ef91a5..bef49e1cf0329 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean @@ -136,14 +136,12 @@ theorem isLeast_opNorm [RingHomIsometric σ₁₂] (f : E →SL[σ₁₂] F) : simp only [setOf_and, setOf_forall] refine isClosed_Ici.inter <| isClosed_iInter fun _ ↦ isClosed_le ?_ ?_ <;> continuity -@[deprecated (since := "2024-02-02")] alias isLeast_op_norm := isLeast_opNorm /-- If one controls the norm of every `A x`, then one controls the norm of `A`. -/ theorem opNorm_le_bound (f : E →SL[σ₁₂] F) {M : ℝ} (hMp : 0 ≤ M) (hM : ∀ x, ‖f x‖ ≤ M * ‖x‖) : ‖f‖ ≤ M := csInf_le bounds_bddBelow ⟨hMp, hM⟩ -@[deprecated (since := "2024-02-02")] alias op_norm_le_bound := opNorm_le_bound /-- If one controls the norm of every `A x`, `‖x‖ ≠ 0`, then one controls the norm of `A`. -/ theorem opNorm_le_bound' (f : E →SL[σ₁₂] F) {M : ℝ} (hMp : 0 ≤ M) @@ -152,13 +150,11 @@ theorem opNorm_le_bound' (f : E →SL[σ₁₂] F) {M : ℝ} (hMp : 0 ≤ M) (ne_or_eq ‖x‖ 0).elim (hM x) fun h => by simp only [h, mul_zero, norm_image_of_norm_zero f f.2 h, le_refl] -@[deprecated (since := "2024-02-02")] alias op_norm_le_bound' := opNorm_le_bound' theorem opNorm_le_of_lipschitz {f : E →SL[σ₁₂] F} {K : ℝ≥0} (hf : LipschitzWith K f) : ‖f‖ ≤ K := f.opNorm_le_bound K.2 fun x => by simpa only [dist_zero_right, f.map_zero] using hf.dist_le_mul x 0 -@[deprecated (since := "2024-02-02")] alias op_norm_le_of_lipschitz := opNorm_le_of_lipschitz theorem opNorm_eq_of_bounds {φ : E →SL[σ₁₂] F} {M : ℝ} (M_nonneg : 0 ≤ M) (h_above : ∀ x, ‖φ x‖ ≤ M * ‖x‖) (h_below : ∀ N ≥ 0, (∀ x, ‖φ x‖ ≤ N * ‖x‖) → M ≤ N) : @@ -167,22 +163,18 @@ theorem opNorm_eq_of_bounds {φ : E →SL[σ₁₂] F} {M : ℝ} (M_nonneg : 0 ((le_csInf_iff ContinuousLinearMap.bounds_bddBelow ⟨M, M_nonneg, h_above⟩).mpr fun N ⟨N_nonneg, hN⟩ => h_below N N_nonneg hN) -@[deprecated (since := "2024-02-02")] alias op_norm_eq_of_bounds := opNorm_eq_of_bounds theorem opNorm_neg (f : E →SL[σ₁₂] F) : ‖-f‖ = ‖f‖ := by simp only [norm_def, neg_apply, norm_neg] -@[deprecated (since := "2024-02-02")] alias op_norm_neg := opNorm_neg theorem opNorm_nonneg (f : E →SL[σ₁₂] F) : 0 ≤ ‖f‖ := Real.sInf_nonneg fun _ ↦ And.left -@[deprecated (since := "2024-02-02")] alias op_norm_nonneg := opNorm_nonneg /-- The norm of the `0` operator is `0`. -/ theorem opNorm_zero : ‖(0 : E →SL[σ₁₂] F)‖ = 0 := le_antisymm (opNorm_le_bound _ le_rfl fun _ ↦ by simp) (opNorm_nonneg _) -@[deprecated (since := "2024-02-02")] alias op_norm_zero := opNorm_zero /-- The norm of the identity is at most `1`. It is in fact `1`, except when the space is trivial where it is `0`. It means that one can not do better than an inequality in general. -/ @@ -197,51 +189,42 @@ variable [RingHomIsometric σ₁₂] [RingHomIsometric σ₂₃] (f g : E →SL[ /-- The fundamental property of the operator norm: `‖f x‖ ≤ ‖f‖ * ‖x‖`. -/ theorem le_opNorm : ‖f x‖ ≤ ‖f‖ * ‖x‖ := (isLeast_opNorm f).1.2 x -@[deprecated (since := "2024-02-02")] alias le_op_norm := le_opNorm theorem dist_le_opNorm (x y : E) : dist (f x) (f y) ≤ ‖f‖ * dist x y := by simp_rw [dist_eq_norm, ← map_sub, f.le_opNorm] -@[deprecated (since := "2024-02-02")] alias dist_le_op_norm := dist_le_opNorm theorem le_of_opNorm_le_of_le {x} {a b : ℝ} (hf : ‖f‖ ≤ a) (hx : ‖x‖ ≤ b) : ‖f x‖ ≤ a * b := (f.le_opNorm x).trans <| by gcongr; exact (opNorm_nonneg f).trans hf -@[deprecated (since := "2024-02-02")] alias le_of_op_norm_le_of_le := le_of_opNorm_le_of_le theorem le_opNorm_of_le {c : ℝ} {x} (h : ‖x‖ ≤ c) : ‖f x‖ ≤ ‖f‖ * c := f.le_of_opNorm_le_of_le le_rfl h -@[deprecated (since := "2024-02-02")] alias le_op_norm_of_le := le_opNorm_of_le theorem le_of_opNorm_le {c : ℝ} (h : ‖f‖ ≤ c) (x : E) : ‖f x‖ ≤ c * ‖x‖ := f.le_of_opNorm_le_of_le h le_rfl -@[deprecated (since := "2024-02-02")] alias le_of_op_norm_le := le_of_opNorm_le theorem opNorm_le_iff {f : E →SL[σ₁₂] F} {M : ℝ} (hMp : 0 ≤ M) : ‖f‖ ≤ M ↔ ∀ x, ‖f x‖ ≤ M * ‖x‖ := ⟨f.le_of_opNorm_le, opNorm_le_bound f hMp⟩ -@[deprecated (since := "2024-02-02")] alias op_norm_le_iff := opNorm_le_iff theorem ratio_le_opNorm : ‖f x‖ / ‖x‖ ≤ ‖f‖ := div_le_of_le_mul₀ (norm_nonneg _) f.opNorm_nonneg (le_opNorm _ _) -@[deprecated (since := "2024-02-02")] alias ratio_le_op_norm := ratio_le_opNorm /-- The image of the unit ball under a continuous linear map is bounded. -/ theorem unit_le_opNorm : ‖x‖ ≤ 1 → ‖f x‖ ≤ ‖f‖ := mul_one ‖f‖ ▸ f.le_opNorm_of_le -@[deprecated (since := "2024-02-02")] alias unit_le_op_norm := unit_le_opNorm theorem opNorm_le_of_shell {f : E →SL[σ₁₂] F} {ε C : ℝ} (ε_pos : 0 < ε) (hC : 0 ≤ C) {c : 𝕜} (hc : 1 < ‖c‖) (hf : ∀ x, ε / ‖c‖ ≤ ‖x‖ → ‖x‖ < ε → ‖f x‖ ≤ C * ‖x‖) : ‖f‖ ≤ C := f.opNorm_le_bound' hC fun _ hx => SemilinearMapClass.bound_of_shell_semi_normed f ε_pos hc hf hx -@[deprecated (since := "2024-02-02")] alias op_norm_le_of_shell := opNorm_le_of_shell theorem opNorm_le_of_ball {f : E →SL[σ₁₂] F} {ε : ℝ} {C : ℝ} (ε_pos : 0 < ε) (hC : 0 ≤ C) (hf : ∀ x ∈ ball (0 : E) ε, ‖f x‖ ≤ C * ‖x‖) : ‖f‖ ≤ C := by @@ -249,14 +232,12 @@ theorem opNorm_le_of_ball {f : E →SL[σ₁₂] F} {ε : ℝ} {C : ℝ} (ε_pos refine opNorm_le_of_shell ε_pos hC hc fun x _ hx => hf x ?_ rwa [ball_zero_eq] -@[deprecated (since := "2024-02-02")] alias op_norm_le_of_ball := opNorm_le_of_ball theorem opNorm_le_of_nhds_zero {f : E →SL[σ₁₂] F} {C : ℝ} (hC : 0 ≤ C) (hf : ∀ᶠ x in 𝓝 (0 : E), ‖f x‖ ≤ C * ‖x‖) : ‖f‖ ≤ C := let ⟨_, ε0, hε⟩ := Metric.eventually_nhds_iff_ball.1 hf opNorm_le_of_ball ε0 hC hε -@[deprecated (since := "2024-02-02")] alias op_norm_le_of_nhds_zero := opNorm_le_of_nhds_zero theorem opNorm_le_of_shell' {f : E →SL[σ₁₂] F} {ε C : ℝ} (ε_pos : 0 < ε) (hC : 0 ≤ C) {c : 𝕜} (hc : ‖c‖ < 1) (hf : ∀ x, ε * ‖c‖ ≤ ‖x‖ → ‖x‖ < ε → ‖f x‖ ≤ C * ‖x‖) : ‖f‖ ≤ C := by @@ -268,7 +249,6 @@ theorem opNorm_le_of_shell' {f : E →SL[σ₁₂] F} {ε C : ℝ} (ε_pos : 0 < refine opNorm_le_of_shell ε_pos hC hc ?_ rwa [norm_inv, div_eq_mul_inv, inv_inv] -@[deprecated (since := "2024-02-02")] alias op_norm_le_of_shell' := opNorm_le_of_shell' /-- For a continuous real linear map `f`, if one controls the norm of every `f x`, `‖x‖ = 1`, then one controls the norm of `f`. -/ @@ -280,14 +260,12 @@ theorem opNorm_le_of_unit_norm [NormedSpace ℝ E] [NormedSpace ℝ F] {f : E rwa [map_smul, norm_smul, norm_inv, norm_norm, ← div_eq_inv_mul, div_le_iff₀] at H₂ exact (norm_nonneg x).lt_of_ne' hx -@[deprecated (since := "2024-02-02")] alias op_norm_le_of_unit_norm := opNorm_le_of_unit_norm /-- The operator norm satisfies the triangle inequality. -/ theorem opNorm_add_le : ‖f + g‖ ≤ ‖f‖ + ‖g‖ := (f + g).opNorm_le_bound (add_nonneg f.opNorm_nonneg g.opNorm_nonneg) fun x => (norm_add_le_of_le (f.le_opNorm x) (g.le_opNorm x)).trans_eq (add_mul _ _ _).symm -@[deprecated (since := "2024-02-02")] alias op_norm_add_le := opNorm_add_le /-- If there is an element with norm different from `0`, then the norm of the identity equals `1`. (Since we are working with seminorms supposing that the space is non-trivial is not enough.) -/ @@ -303,7 +281,6 @@ theorem opNorm_smul_le {𝕜' : Type*} [NormedField 𝕜'] [NormedSpace 𝕜' F] rw [smul_apply, norm_smul, mul_assoc] exact mul_le_mul_of_nonneg_left (le_opNorm _ _) (norm_nonneg _) -@[deprecated (since := "2024-02-02")] alias op_norm_smul_le := opNorm_smul_le /-- Operator seminorm on the space of continuous (semi)linear maps, as `Seminorm`. @@ -352,7 +329,6 @@ theorem opNorm_comp_le (f : E →SL[σ₁₂] F) : ‖h.comp f‖ ≤ ‖h‖ * rw [mul_assoc] exact h.le_opNorm_of_le (f.le_opNorm x)⟩ -@[deprecated (since := "2024-02-02")] alias op_norm_comp_le := opNorm_comp_le /-- Continuous linear maps form a seminormed ring with respect to the operator norm. -/ instance toSemiNormedRing : SeminormedRing (E →L[𝕜] E) := @@ -378,7 +354,6 @@ theorem opNorm_subsingleton [Subsingleton E] : ‖f‖ = 0 := by intro x simp [Subsingleton.elim x 0] -@[deprecated (since := "2024-02-02")] alias op_norm_subsingleton := opNorm_subsingleton end OpNorm diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean index 9cb3ac62feea8..8193941231cfa 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean @@ -55,7 +55,6 @@ theorem opNorm_ext [RingHomIsometric σ₁₃] (f : E →SL[σ₁₂] F) (g : E rw [← h z] exact h₂ z -@[deprecated (since := "2024-02-02")] alias op_norm_ext := opNorm_ext variable [RingHomIsometric σ₂₃] @@ -63,20 +62,17 @@ theorem opNorm_le_bound₂ (f : E →SL[σ₁₃] F →SL[σ₂₃] G) {C : ℝ} (hC : ∀ x y, ‖f x y‖ ≤ C * ‖x‖ * ‖y‖) : ‖f‖ ≤ C := f.opNorm_le_bound h0 fun x => (f x).opNorm_le_bound (mul_nonneg h0 (norm_nonneg _)) <| hC x -@[deprecated (since := "2024-02-02")] alias op_norm_le_bound₂ := opNorm_le_bound₂ theorem le_opNorm₂ [RingHomIsometric σ₁₃] (f : E →SL[σ₁₃] F →SL[σ₂₃] G) (x : E) (y : F) : ‖f x y‖ ≤ ‖f‖ * ‖x‖ * ‖y‖ := (f x).le_of_opNorm_le (f.le_opNorm x) y -@[deprecated (since := "2024-02-02")] alias le_op_norm₂ := le_opNorm₂ theorem le_of_opNorm₂_le_of_le [RingHomIsometric σ₁₃] (f : E →SL[σ₁₃] F →SL[σ₂₃] G) {x : E} {y : F} {a b c : ℝ} (hf : ‖f‖ ≤ a) (hx : ‖x‖ ≤ b) (hy : ‖y‖ ≤ c) : ‖f x y‖ ≤ a * b * c := (f x).le_of_opNorm_le_of_le (f.le_of_opNorm_le_of_le hf hx) hy -@[deprecated (since := "2024-02-02")] alias le_of_op_norm₂_le_of_le := le_of_opNorm₂_le_of_le end OpNorm @@ -172,7 +168,6 @@ theorem flip_flip (f : E →SL[σ₁₃] F →SL[σ₂₃] G) : f.flip.flip = f theorem opNorm_flip (f : E →SL[σ₁₃] F →SL[σ₂₃] G) : ‖f.flip‖ = ‖f‖ := le_antisymm (by simpa only [flip_flip] using le_norm_flip f.flip) (le_norm_flip f) -@[deprecated (since := "2024-02-02")] alias op_norm_flip := opNorm_flip @[simp] theorem flip_add (f g : E →SL[σ₁₃] F →SL[σ₂₃] G) : (f + g).flip = f.flip + g.flip := diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean index 4043b6833c43d..39bbe337bb7d7 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean @@ -247,7 +247,6 @@ theorem opNorm_extend_le : _ ≤ ‖f‖ * (N * ‖e x‖) := mul_le_mul_of_nonneg_left (h_e x) (norm_nonneg _) _ ≤ N * ‖f‖ * ‖e x‖ := by rw [mul_comm ↑N ‖f‖, mul_assoc] -@[deprecated (since := "2024-02-02")] alias op_norm_extend_le := opNorm_extend_le end diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean index c45a659de2c23..147350c036a57 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean @@ -45,12 +45,10 @@ theorem mul_apply' (x y : 𝕜') : mul 𝕜 𝕜' x y = x * y := theorem opNorm_mul_apply_le (x : 𝕜') : ‖mul 𝕜 𝕜' x‖ ≤ ‖x‖ := opNorm_le_bound _ (norm_nonneg x) (norm_mul_le x) -@[deprecated (since := "2024-02-02")] alias op_norm_mul_apply_le := opNorm_mul_apply_le theorem opNorm_mul_le : ‖mul 𝕜 𝕜'‖ ≤ 1 := LinearMap.mkContinuous₂_norm_le _ zero_le_one _ -@[deprecated (since := "2024-02-02")] alias op_norm_mul_le := opNorm_mul_le /-- Multiplication on the left in a non-unital normed algebra `𝕜'` as a non-unital algebra homomorphism into the algebra of *continuous* linear maps. This is the left regular representation @@ -87,22 +85,14 @@ theorem opNorm_mulLeftRight_apply_apply_le (x y : 𝕜') : ‖mulLeftRight 𝕜 (opNorm_le_bound _ (norm_nonneg _) fun _ => (norm_mul_le _ _).trans_eq (mul_comm _ _)) (norm_nonneg _) (norm_nonneg _) -@[deprecated (since := "2024-02-02")] -alias op_norm_mulLeftRight_apply_apply_le := - opNorm_mulLeftRight_apply_apply_le - theorem opNorm_mulLeftRight_apply_le (x : 𝕜') : ‖mulLeftRight 𝕜 𝕜' x‖ ≤ ‖x‖ := opNorm_le_bound _ (norm_nonneg x) (opNorm_mulLeftRight_apply_apply_le 𝕜 𝕜' x) -@[deprecated (since := "2024-02-02")] -alias op_norm_mulLeftRight_apply_le := opNorm_mulLeftRight_apply_le - set_option maxSynthPendingDepth 2 in theorem opNorm_mulLeftRight_le : ‖mulLeftRight 𝕜 𝕜'‖ ≤ 1 := opNorm_le_bound _ zero_le_one fun x => (one_mul ‖x‖).symm ▸ opNorm_mulLeftRight_apply_le 𝕜 𝕜' x -@[deprecated (since := "2024-02-02")] alias op_norm_mulLeftRight_le := 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 @@ -133,13 +123,11 @@ lemma isometry_mul : Isometry (mul 𝕜 𝕜') := lemma opNorm_mul_apply (x : 𝕜') : ‖mul 𝕜 𝕜' x‖ = ‖x‖ := (AddMonoidHomClass.isometry_iff_norm (mul 𝕜 𝕜')).mp (isometry_mul 𝕜 𝕜') x -@[deprecated (since := "2024-02-02")] alias op_norm_mul_apply := opNorm_mul_apply @[simp] lemma opNNNorm_mul_apply (x : 𝕜') : ‖mul 𝕜 𝕜' x‖₊ = ‖x‖₊ := Subtype.ext <| opNorm_mul_apply 𝕜 𝕜' x -@[deprecated (since := "2024-02-02")] alias op_nnnorm_mul_apply := opNNNorm_mul_apply /-- Multiplication in a normed algebra as a linear isometry to the space of continuous linear maps. -/ @@ -210,7 +198,6 @@ variable {𝕜} theorem opNorm_lsmul_apply_le (x : 𝕜') : ‖(lsmul 𝕜 𝕜' x : E →L[𝕜] E)‖ ≤ ‖x‖ := ContinuousLinearMap.opNorm_le_bound _ (norm_nonneg x) fun y => norm_smul_le x y -@[deprecated (since := "2024-02-02")] alias op_norm_lsmul_apply_le := opNorm_lsmul_apply_le /-- The norm of `lsmul` is at most 1 in any semi-normed group. -/ theorem opNorm_lsmul_le : ‖(lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] E →L[𝕜] E)‖ ≤ 1 := by @@ -218,7 +205,6 @@ theorem opNorm_lsmul_le : ‖(lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] E →L[𝕜] E simp_rw [one_mul] exact opNorm_lsmul_apply_le _ -@[deprecated (since := "2024-02-02")] alias op_norm_lsmul_le := opNorm_lsmul_le end SMulLinear @@ -242,13 +228,11 @@ variable [SMulCommClass 𝕜 𝕜' 𝕜'] [RegularNormedAlgebra 𝕜 𝕜'] [Non theorem opNorm_mul : ‖mul 𝕜 𝕜'‖ = 1 := (mulₗᵢ 𝕜 𝕜').norm_toContinuousLinearMap -@[deprecated (since := "2024-02-02")] alias op_norm_mul := opNorm_mul @[simp] theorem opNNNorm_mul : ‖mul 𝕜 𝕜'‖₊ = 1 := Subtype.ext <| opNorm_mul 𝕜 𝕜' -@[deprecated (since := "2024-02-02")] alias op_nnnorm_mul := opNNNorm_mul end @@ -267,7 +251,6 @@ theorem opNorm_lsmul [NormedField 𝕜'] [NormedAlgebra 𝕜 𝕜'] [NormedSpace refine le_of_mul_le_mul_right ?_ (norm_pos_iff.mpr hy) simp_rw [one_mul, this] -@[deprecated (since := "2024-02-02")] alias op_norm_lsmul := opNorm_lsmul end ContinuousLinearMap diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean index 51f3b1705bfc9..018b01d78582b 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean @@ -56,14 +56,12 @@ theorem nnnorm_def (f : E →SL[σ₁₂] F) : ‖f‖₊ = sInf { c | ∀ x, theorem opNNNorm_le_bound (f : E →SL[σ₁₂] F) (M : ℝ≥0) (hM : ∀ x, ‖f x‖₊ ≤ M * ‖x‖₊) : ‖f‖₊ ≤ M := opNorm_le_bound f (zero_le M) hM -@[deprecated (since := "2024-02-02")] alias op_nnnorm_le_bound := opNNNorm_le_bound /-- If one controls the norm of every `A x`, `‖x‖₊ ≠ 0`, then one controls the norm of `A`. -/ theorem opNNNorm_le_bound' (f : E →SL[σ₁₂] F) (M : ℝ≥0) (hM : ∀ x, ‖x‖₊ ≠ 0 → ‖f x‖₊ ≤ M * ‖x‖₊) : ‖f‖₊ ≤ M := opNorm_le_bound' f (zero_le M) fun x hx => hM x <| by rwa [← NNReal.coe_ne_zero] -@[deprecated (since := "2024-02-02")] alias op_nnnorm_le_bound' := opNNNorm_le_bound' /-- For a continuous real linear map `f`, if one controls the norm of every `f x`, `‖x‖₊ = 1`, then one controls the norm of `f`. -/ @@ -71,45 +69,35 @@ theorem opNNNorm_le_of_unit_nnnorm [NormedSpace ℝ E] [NormedSpace ℝ F] {f : (hf : ∀ x, ‖x‖₊ = 1 → ‖f x‖₊ ≤ C) : ‖f‖₊ ≤ C := opNorm_le_of_unit_norm C.coe_nonneg fun x hx => hf x <| by rwa [← NNReal.coe_eq_one] -@[deprecated (since := "2024-02-02")] -alias op_nnnorm_le_of_unit_nnnorm := opNNNorm_le_of_unit_nnnorm - theorem opNNNorm_le_of_lipschitz {f : E →SL[σ₁₂] F} {K : ℝ≥0} (hf : LipschitzWith K f) : ‖f‖₊ ≤ K := opNorm_le_of_lipschitz hf -@[deprecated (since := "2024-02-02")] alias op_nnnorm_le_of_lipschitz := opNNNorm_le_of_lipschitz theorem opNNNorm_eq_of_bounds {φ : E →SL[σ₁₂] F} (M : ℝ≥0) (h_above : ∀ x, ‖φ x‖₊ ≤ M * ‖x‖₊) (h_below : ∀ N, (∀ x, ‖φ x‖₊ ≤ N * ‖x‖₊) → M ≤ N) : ‖φ‖₊ = M := Subtype.ext <| opNorm_eq_of_bounds (zero_le M) h_above <| Subtype.forall'.mpr h_below -@[deprecated (since := "2024-02-02")] alias op_nnnorm_eq_of_bounds := opNNNorm_eq_of_bounds theorem opNNNorm_le_iff {f : E →SL[σ₁₂] F} {C : ℝ≥0} : ‖f‖₊ ≤ C ↔ ∀ x, ‖f x‖₊ ≤ C * ‖x‖₊ := opNorm_le_iff C.2 -@[deprecated (since := "2024-02-02")] alias op_nnnorm_le_iff := opNNNorm_le_iff theorem isLeast_opNNNorm : IsLeast {C : ℝ≥0 | ∀ x, ‖f x‖₊ ≤ C * ‖x‖₊} ‖f‖₊ := by simpa only [← opNNNorm_le_iff] using isLeast_Ici -@[deprecated (since := "2024-02-02")] alias isLeast_op_nnnorm := isLeast_opNNNorm theorem opNNNorm_comp_le [RingHomIsometric σ₁₃] (f : E →SL[σ₁₂] F) : ‖h.comp f‖₊ ≤ ‖h‖₊ * ‖f‖₊ := opNorm_comp_le h f -@[deprecated (since := "2024-02-02")] alias op_nnnorm_comp_le := opNNNorm_comp_le theorem le_opNNNorm : ‖f x‖₊ ≤ ‖f‖₊ * ‖x‖₊ := f.le_opNorm x -@[deprecated (since := "2024-02-02")] alias le_op_nnnorm := le_opNNNorm theorem nndist_le_opNNNorm (x y : E) : nndist (f x) (f y) ≤ ‖f‖₊ * nndist x y := dist_le_opNorm f x y -@[deprecated (since := "2024-02-02")] alias nndist_le_op_nnnorm := nndist_le_opNNNorm /-- continuous linear maps are Lipschitz continuous. -/ theorem lipschitz : LipschitzWith ‖f‖₊ f := @@ -131,17 +119,11 @@ theorem exists_mul_lt_apply_of_lt_opNNNorm (f : E →SL[σ₁₂] F) {r : ℝ≥ not_mem_of_lt_csInf (nnnorm_def f ▸ hr : r < sInf { c : ℝ≥0 | ∀ x, ‖f x‖₊ ≤ c * ‖x‖₊ }) (OrderBot.bddBelow _) -@[deprecated (since := "2024-02-02")] -alias exists_mul_lt_apply_of_lt_op_nnnorm := exists_mul_lt_apply_of_lt_opNNNorm - theorem exists_mul_lt_of_lt_opNorm (f : E →SL[σ₁₂] F) {r : ℝ} (hr₀ : 0 ≤ r) (hr : r < ‖f‖) : ∃ x, r * ‖x‖ < ‖f x‖ := by lift r to ℝ≥0 using hr₀ exact f.exists_mul_lt_apply_of_lt_opNNNorm hr -@[deprecated (since := "2024-02-02")] -alias exists_mul_lt_of_lt_op_norm := exists_mul_lt_of_lt_opNorm - theorem exists_lt_apply_of_lt_opNNNorm {𝕜 𝕜₂ E F : Type*} [NormedAddCommGroup E] [SeminormedAddCommGroup F] [DenselyNormedField 𝕜] [NontriviallyNormedField 𝕜₂] {σ₁₂ : 𝕜 →+* 𝕜₂} [NormedSpace 𝕜 E] [NormedSpace 𝕜₂ F] [RingHomIsometric σ₁₂] (f : E →SL[σ₁₂] F) {r : ℝ≥0} @@ -158,9 +140,6 @@ theorem exists_lt_apply_of_lt_opNNNorm {𝕜 𝕜₂ E F : Type*} [NormedAddComm have : ‖σ₁₂ k‖₊ = ‖k‖₊ := Subtype.ext RingHomIsometric.is_iso rwa [map_smulₛₗ f, nnnorm_smul, ← div_lt_iff₀ hfy.bot_lt, div_eq_mul_inv, this] -@[deprecated (since := "2024-02-02")] -alias exists_lt_apply_of_lt_op_nnnorm := exists_lt_apply_of_lt_opNNNorm - theorem exists_lt_apply_of_lt_opNorm {𝕜 𝕜₂ E F : Type*} [NormedAddCommGroup E] [SeminormedAddCommGroup F] [DenselyNormedField 𝕜] [NontriviallyNormedField 𝕜₂] {σ₁₂ : 𝕜 →+* 𝕜₂} [NormedSpace 𝕜 E] [NormedSpace 𝕜₂ F] [RingHomIsometric σ₁₂] (f : E →SL[σ₁₂] F) {r : ℝ} @@ -170,9 +149,6 @@ theorem exists_lt_apply_of_lt_opNorm {𝕜 𝕜₂ E F : Type*} [NormedAddCommGr · lift r to ℝ≥0 using not_lt.1 hr₀ exact f.exists_lt_apply_of_lt_opNNNorm hr -@[deprecated (since := "2024-02-02")] -alias exists_lt_apply_of_lt_op_norm := exists_lt_apply_of_lt_opNorm - theorem sSup_unit_ball_eq_nnnorm {𝕜 𝕜₂ E F : Type*} [NormedAddCommGroup E] [SeminormedAddCommGroup F] [DenselyNormedField 𝕜] [NontriviallyNormedField 𝕜₂] {σ₁₂ : 𝕜 →+* 𝕜₂} [NormedSpace 𝕜 E] [NormedSpace 𝕜₂ F] [RingHomIsometric σ₁₂] (f : E →SL[σ₁₂] F) : diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean index b08ba685f3320..f9b408f94f338 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean @@ -102,7 +102,6 @@ theorem opNorm_zero_iff [RingHomIsometric σ₁₂] : ‖f‖ = 0 ↔ f = 0 := rintro rfl exact opNorm_zero) -@[deprecated (since := "2024-02-02")] alias op_norm_zero_iff := opNorm_zero_iff /-- If a normed space is non-trivial, then the norm of the identity equals `1`. -/ @[simp] @@ -208,9 +207,6 @@ theorem opNorm_comp_linearIsometryEquiv (f : F →SL[σ₂₃] G) (g : F' ≃ₛ haveI := g.symm.surjective.nontrivial simp [g.symm.toLinearIsometry.norm_toContinuousLinearMap] -@[deprecated (since := "2024-02-02")] -alias op_norm_comp_linearIsometryEquiv := opNorm_comp_linearIsometryEquiv - @[simp] theorem norm_smulRightL (c : E →L[𝕜] 𝕜) [Nontrivial Fₗ] : ‖smulRightL 𝕜 E Fₗ c‖ = ‖c‖ := ContinuousLinearMap.homothety_norm _ c.norm_smulRight_apply diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean index 51e531175bf03..8d1a93937e77e 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean @@ -52,13 +52,11 @@ theorem opNorm_prod (f : E →L[𝕜] F) (g : E →L[𝕜] G) : ‖f.prod g‖ = (opNorm_le_bound _ (norm_nonneg _) fun x => (le_max_right _ _).trans ((f.prod g).le_opNorm x)) -@[deprecated (since := "2024-02-02")] alias op_norm_prod := opNorm_prod @[simp] theorem opNNNorm_prod (f : E →L[𝕜] F) (g : E →L[𝕜] G) : ‖f.prod g‖₊ = ‖(f, g)‖₊ := Subtype.ext <| opNorm_prod f g -@[deprecated (since := "2024-02-02")] alias op_nnnorm_prod := opNNNorm_prod /-- `ContinuousLinearMap.prod` as a `LinearIsometryEquiv`. -/ def prodₗᵢ (R : Type*) [Semiring R] [Module R F] [Module R G] [ContinuousConstSMul R F] diff --git a/Mathlib/Analysis/NormedSpace/RCLike.lean b/Mathlib/Analysis/NormedSpace/RCLike.lean index f578908f625cc..a918890d4c64b 100644 --- a/Mathlib/Analysis/NormedSpace/RCLike.lean +++ b/Mathlib/Analysis/NormedSpace/RCLike.lean @@ -85,10 +85,6 @@ theorem ContinuousLinearMap.opNorm_bound_of_ball_bound {r : ℝ} (r_pos : 0 < r) apply LinearMap.bound_of_ball_bound' r_pos exact fun z hz => h z hz -@[deprecated (since := "2024-02-02")] -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) : diff --git a/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean b/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean index 89db8290e3687..ee0f874953f99 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean @@ -210,9 +210,6 @@ theorem _root_.fourierIntegral_gaussian (hb : 0 < b.re) (t : ℂ) : rw [integral_cexp_quadratic (show (-b).re < 0 by rwa [neg_re, neg_lt_zero]), neg_neg, zero_sub, mul_neg, div_neg, neg_neg, mul_pow, I_sq, neg_one_mul, mul_comm] -@[deprecated (since := "2024-02-21")] -alias _root_.fourier_transform_gaussian := fourierIntegral_gaussian - theorem _root_.fourierIntegral_gaussian_pi' (hb : 0 < b.re) (c : ℂ) : (𝓕 fun x : ℝ => cexp (-π * b * x ^ 2 + 2 * π * c * x)) = fun t : ℝ => 1 / b ^ (1 / 2 : ℂ) * cexp (-π / b * (t + I * c) ^ 2) := by @@ -233,17 +230,11 @@ theorem _root_.fourierIntegral_gaussian_pi' (hb : 0 < b.re) (c : ℂ) : simp only [I_sq] ring -@[deprecated (since := "2024-02-21")] -alias _root_.fourier_transform_gaussian_pi' := _root_.fourierIntegral_gaussian_pi' - theorem _root_.fourierIntegral_gaussian_pi (hb : 0 < b.re) : (𝓕 fun (x : ℝ) ↦ cexp (-π * b * x ^ 2)) = fun t : ℝ ↦ 1 / b ^ (1 / 2 : ℂ) * cexp (-π / b * t ^ 2) := by simpa only [mul_zero, zero_mul, add_zero] using fourierIntegral_gaussian_pi' hb 0 -@[deprecated (since := "2024-02-21")] -alias root_.fourier_transform_gaussian_pi := _root_.fourierIntegral_gaussian_pi - section InnerProductSpace variable {V : Type*} [NormedAddCommGroup V] [InnerProductSpace ℝ V] [FiniteDimensional ℝ V] diff --git a/Mathlib/CategoryTheory/Filtered/Basic.lean b/Mathlib/CategoryTheory/Filtered/Basic.lean index 584411d3326a4..e356e98a41ca6 100644 --- a/Mathlib/CategoryTheory/Filtered/Basic.lean +++ b/Mathlib/CategoryTheory/Filtered/Basic.lean @@ -829,9 +829,6 @@ theorem of_isInitial {X : C} (h : IsInitial X) : IsCofiltered C := instance (priority := 100) of_hasInitial [HasInitial C] : IsCofiltered C := of_isInitial _ initialIsInitial -@[deprecated (since := "2024-03-11")] -alias _root_.CategoryTheory.cofiltered_of_hasFiniteLimits := of_hasFiniteLimits - /-- For every universe `w`, `C` is filtered if and only if every finite diagram in `C` with shape in `w` admits a cocone. -/ theorem iff_cone_nonempty : IsCofiltered C ↔ diff --git a/Mathlib/CategoryTheory/IsConnected.lean b/Mathlib/CategoryTheory/IsConnected.lean index 9f27285ea78dc..ebe6c3c30a930 100644 --- a/Mathlib/CategoryTheory/IsConnected.lean +++ b/Mathlib/CategoryTheory/IsConnected.lean @@ -373,7 +373,6 @@ theorem isPreconnected_zigzag [IsPreconnected J] (j₁ j₂ : J) : Zigzag j₁ j equiv_relation _ zigzag_equivalence (fun f => Relation.ReflTransGen.single (Or.inl (Nonempty.intro f))) _ _ -@[deprecated (since := "2024-02-19")] alias isConnected_zigzag := isPreconnected_zigzag theorem zigzag_isPreconnected (h : ∀ j₁ j₂ : J, Zigzag j₁ j₂) : IsPreconnected J := by apply IsPreconnected.of_constant_of_preserves_morphisms @@ -454,7 +453,4 @@ theorem nonempty_hom_of_preconnected_groupoid {G} [Groupoid G] [IsPreconnected G attribute [instance] nonempty_hom_of_preconnected_groupoid -@[deprecated (since := "2024-02-19")] -alias nonempty_hom_of_connected_groupoid := nonempty_hom_of_preconnected_groupoid - end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/FullSubcategory.lean b/Mathlib/CategoryTheory/Limits/FullSubcategory.lean index 18abbde52961c..11c03c8e182d8 100644 --- a/Mathlib/CategoryTheory/Limits/FullSubcategory.lean +++ b/Mathlib/CategoryTheory/Limits/FullSubcategory.lean @@ -113,8 +113,6 @@ theorem hasLimit_of_closedUnderLimits (h : ClosedUnderLimitsOfShape J P) have : CreatesLimit F (fullSubcategoryInclusion P) := createsLimitFullSubcategoryInclusionOfClosed h F hasLimit_of_created F (fullSubcategoryInclusion P) -@[deprecated (since := "2024-03-23")] -alias hasLimit_of_closed_under_limits := hasLimit_of_closedUnderLimits theorem hasLimitsOfShape_of_closedUnderLimits (h : ClosedUnderLimitsOfShape J P) [HasLimitsOfShape J C] : HasLimitsOfShape J (FullSubcategory P) := @@ -136,14 +134,10 @@ theorem hasColimit_of_closedUnderColimits (h : ClosedUnderColimitsOfShape J P) have : CreatesColimit F (fullSubcategoryInclusion P) := createsColimitFullSubcategoryInclusionOfClosed h F hasColimit_of_created F (fullSubcategoryInclusion P) -@[deprecated (since := "2024-03-23")] -alias hasColimit_of_closed_under_colimits := hasColimit_of_closedUnderColimits theorem hasColimitsOfShape_of_closedUnderColimits (h : ClosedUnderColimitsOfShape J P) [HasColimitsOfShape J C] : HasColimitsOfShape J (FullSubcategory P) := { has_colimit := fun F => hasColimit_of_closedUnderColimits h F } -@[deprecated (since := "2024-03-23")] -alias hasColimitsOfShape_of_closed_under_colimits := hasColimitsOfShape_of_closedUnderColimits end diff --git a/Mathlib/CategoryTheory/Limits/Opposites.lean b/Mathlib/CategoryTheory/Limits/Opposites.lean index 76a2af3431f75..7c040f35246ea 100644 --- a/Mathlib/CategoryTheory/Limits/Opposites.lean +++ b/Mathlib/CategoryTheory/Limits/Opposites.lean @@ -32,10 +32,6 @@ namespace CategoryTheory.Limits variable {C : Type u₁} [Category.{v₁} C] variable {J : Type u₂} [Category.{v₂} J] -@[deprecated (since := "2024-03-26")] alias isLimitCoconeOp := IsColimit.op -@[deprecated (since := "2024-03-26")] alias isColimitConeOp := IsLimit.op -@[deprecated (since := "2024-03-26")] alias isLimitCoconeUnop := IsColimit.unop -@[deprecated (since := "2024-03-26")] alias isColimitConeUnop := IsLimit.unop /-- Turn a colimit for `F : J ⥤ Cᵒᵖ` into a limit for `F.leftOp : Jᵒᵖ ⥤ C`. -/ @[simps] diff --git a/Mathlib/Combinatorics/SimpleGraph/Coloring.lean b/Mathlib/Combinatorics/SimpleGraph/Coloring.lean index 9d98315f91068..7f5fe6d81b937 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Coloring.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Coloring.lean @@ -265,10 +265,6 @@ theorem chromaticNumber_le_iff_colorable {n : ℕ} : G.chromaticNumber ≤ n ↔ rw [Set.mem_setOf_eq] at this exact this.mono h -@[deprecated Colorable.chromaticNumber_le (since := "2024-03-21")] -theorem chromaticNumber_le_card [Fintype α] (C : G.Coloring α) : - G.chromaticNumber ≤ Fintype.card α := C.colorable.chromaticNumber_le - theorem colorable_chromaticNumber {m : ℕ} (hc : G.Colorable m) : G.Colorable (ENat.toNat G.chromaticNumber) := by classical diff --git a/Mathlib/Data/Complex/Abs.lean b/Mathlib/Data/Complex/Abs.lean index 4735e0966f985..7531c38a54ff3 100644 --- a/Mathlib/Data/Complex/Abs.lean +++ b/Mathlib/Data/Complex/Abs.lean @@ -198,9 +198,6 @@ theorem abs_im_div_abs_le_one (z : ℂ) : |z.im / Complex.abs z| ≤ 1 := @[simp, norm_cast] lemma abs_intCast (n : ℤ) : abs n = |↑n| := by rw [← ofReal_intCast, abs_ofReal] -@[deprecated "No deprecation message was provided." (since := "2024-02-14")] -lemma int_cast_abs (n : ℤ) : |↑n| = Complex.abs n := (abs_intCast _).symm - theorem normSq_eq_abs (x : ℂ) : normSq x = (Complex.abs x) ^ 2 := by simp [abs, sq, abs_def, Real.mul_self_sqrt (normSq_nonneg _)] diff --git a/Mathlib/Data/Countable/Small.lean b/Mathlib/Data/Countable/Small.lean index 72387f0c4a096..b48f61815b730 100644 --- a/Mathlib/Data/Countable/Small.lean +++ b/Mathlib/Data/Countable/Small.lean @@ -18,9 +18,3 @@ universe w v instance (priority := 100) Countable.toSmall (α : Type v) [Countable α] : Small.{w} α := let ⟨_, hf⟩ := exists_injective_nat α small_of_injective hf - -@[deprecated (since := "2024-03-20"), nolint defLemma] -alias small_of_countable := Countable.toSmall - -@[deprecated (since := "2024-03-20"), nolint defLemma] -alias small_of_fintype := Countable.toSmall diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 6056e283f70cc..148cf433376f6 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -79,8 +79,6 @@ 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 @@ -151,10 +149,6 @@ section coe theorem val_eq_val (a b : Fin n) : (a : ℕ) = b ↔ a = b := Fin.ext_iff.symm -@[deprecated Fin.ext_iff (since := "2024-02-20")] -theorem eq_iff_veq (a b : Fin n) : a = b ↔ a.1 = b.1 := - Fin.ext_iff - theorem ne_iff_vne (a b : Fin n) : a ≠ b ↔ a.1 ≠ b.1 := Fin.ext_iff.not diff --git a/Mathlib/Data/Finset/Card.lean b/Mathlib/Data/Finset/Card.lean index 648f621a5abac..3d203a0107b42 100644 --- a/Mathlib/Data/Finset/Card.lean +++ b/Mathlib/Data/Finset/Card.lean @@ -500,8 +500,6 @@ lemma card_union_eq_card_add_card : #(s ∪ t) = #s + #t ↔ Disjoint s t := by @[simp] alias ⟨_, card_union_of_disjoint⟩ := card_union_eq_card_add_card -@[deprecated (since := "2024-02-09")] alias card_union_eq := card_union_of_disjoint -@[deprecated (since := "2024-02-09")] alias card_disjoint_union := card_union_of_disjoint lemma cast_card_inter [AddGroupWithOne R] : (#(s ∩ t) : R) = #s + #t - #(s ∪ t) := by @@ -676,9 +674,6 @@ theorem one_lt_card_iff_nontrivial : 1 < #s ↔ s.Nontrivial := by rw [← not_iff_not, not_lt, Finset.Nontrivial, ← Set.nontrivial_coe_sort, not_nontrivial_iff_subsingleton, card_le_one_iff_subsingleton_coe, coe_sort_coe] -@[deprecated (since := "2024-02-05")] -alias one_lt_card_iff_nontrivial_coe := one_lt_card_iff_nontrivial - theorem exists_ne_of_one_lt_card (hs : 1 < #s) (a : α) : ∃ b, b ∈ s ∧ b ≠ a := by obtain ⟨x, hx, y, hy, hxy⟩ := Finset.one_lt_card.mp hs by_cases ha : y = a diff --git a/Mathlib/Data/Finset/Empty.lean b/Mathlib/Data/Finset/Empty.lean index 5916c7d3dc88a..60a6034dfb73c 100644 --- a/Mathlib/Data/Finset/Empty.lean +++ b/Mathlib/Data/Finset/Empty.lean @@ -67,7 +67,6 @@ alias ⟨_, Nonempty.coe_sort⟩ := nonempty_coe_sort theorem Nonempty.exists_mem {s : Finset α} (h : s.Nonempty) : ∃ x : α, x ∈ s := h -@[deprecated (since := "2024-03-23")] alias Nonempty.bex := Nonempty.exists_mem theorem Nonempty.mono {s t : Finset α} (hst : s ⊆ t) (hs : s.Nonempty) : t.Nonempty := Set.Nonempty.mono hst hs diff --git a/Mathlib/Data/Finset/Lattice/Basic.lean b/Mathlib/Data/Finset/Lattice/Basic.lean index 14fc97ad84371..79aeaeff8285a 100644 --- a/Mathlib/Data/Finset/Lattice/Basic.lean +++ b/Mathlib/Data/Finset/Lattice/Basic.lean @@ -295,10 +295,6 @@ theorem union_inter_distrib_left (s t u : Finset α) : s ∪ t ∩ u = (s ∪ t) theorem inter_union_distrib_right (s t u : Finset α) : s ∩ t ∪ u = (s ∪ u) ∩ (t ∪ u) := sup_inf_right _ _ _ -@[deprecated (since := "2024-03-22")] alias inter_distrib_left := inter_union_distrib_left -@[deprecated (since := "2024-03-22")] alias inter_distrib_right := union_inter_distrib_right -@[deprecated (since := "2024-03-22")] alias union_distrib_left := union_inter_distrib_left -@[deprecated (since := "2024-03-22")] alias union_distrib_right := inter_union_distrib_right theorem union_union_distrib_left (s t u : Finset α) : s ∪ (t ∪ u) = s ∪ t ∪ (s ∪ u) := sup_sup_distrib_left _ _ _ diff --git a/Mathlib/Data/Int/Defs.lean b/Mathlib/Data/Int/Defs.lean index 3addc51ed7381..732c14aa0df2c 100644 --- a/Mathlib/Data/Int/Defs.lean +++ b/Mathlib/Data/Int/Defs.lean @@ -91,9 +91,6 @@ protected theorem ofNat_add_one_out (n : ℕ) : ↑n + (1 : ℤ) = ↑(succ n) : @[simp] lemma ofNat_eq_natCast (n : ℕ) : Int.ofNat n = n := rfl -@[deprecated ofNat_eq_natCast (since := "2024-03-24")] -protected lemma natCast_eq_ofNat (n : ℕ) : ↑n = Int.ofNat n := rfl - @[norm_cast] lemma natCast_inj {m n : ℕ} : (m : ℤ) = (n : ℤ) ↔ m = n := ofNat_inj @[simp, norm_cast] lemma natAbs_cast (n : ℕ) : natAbs ↑n = n := rfl diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 4c99fc84425e5..311d98d7e2fcd 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -73,7 +73,6 @@ theorem _root_.Decidable.List.eq_or_ne_mem_of_mem [DecidableEq α] lemma mem_pair {a b c : α} : a ∈ [b, c] ↔ a = b ∨ a = c := by rw [mem_cons, mem_singleton] -@[deprecated (since := "2024-03-23")] alias mem_split := append_of_mem -- The simpNF linter says that the LHS can be simplified via `List.mem_map`. -- However this is a higher priority lemma. @@ -872,11 +871,6 @@ theorem flatMap_pure_eq_map (f : α → β) (l : List α) : l.flatMap (pure ∘ @[deprecated (since := "2024-10-16")] alias bind_pure_eq_map := flatMap_pure_eq_map -set_option linter.deprecated false in -@[deprecated flatMap_pure_eq_map (since := "2024-03-24")] -theorem bind_ret_eq_map (f : α → β) (l : List α) : l.bind (List.ret ∘ f) = map f l := - bind_pure_eq_map f l - theorem flatMap_congr {l : List α} {f g : α → List β} (h : ∀ x ∈ l, f x = g x) : List.flatMap l f = List.flatMap l g := (congr_arg List.flatten <| map_congr_left h : _) diff --git a/Mathlib/Data/List/Defs.lean b/Mathlib/Data/List/Defs.lean index b340e9b270a42..9718d8cc7eb2c 100644 --- a/Mathlib/Data/List/Defs.lean +++ b/Mathlib/Data/List/Defs.lean @@ -504,8 +504,6 @@ alias ⟨eq_or_mem_of_mem_cons, _⟩ := mem_cons theorem not_exists_mem_nil (p : α → Prop) : ¬∃ x ∈ @nil α, p x := fun ⟨_, hx, _⟩ => List.not_mem_nil _ hx -@[deprecated (since := "2024-03-23")] alias not_bex_nil := not_exists_mem_nil -@[deprecated (since := "2024-03-23")] alias bex_cons := exists_mem_cons @[deprecated (since := "2024-08-10")] alias length_le_of_sublist := Sublist.length_le diff --git a/Mathlib/Data/List/Infix.lean b/Mathlib/Data/List/Infix.lean index ddab47721a8c6..942f206a4afa8 100644 --- a/Mathlib/Data/List/Infix.lean +++ b/Mathlib/Data/List/Infix.lean @@ -106,7 +106,6 @@ instance decidableInfix [DecidableEq α] : ∀ l₁ l₂ : List α, Decidable (l theorem cons_prefix_iff : a :: l₁ <+: b :: l₂ ↔ a = b ∧ l₁ <+: l₂ := by simp -@[deprecated (since := "2024-03-26")] alias IsPrefix.filter_map := IsPrefix.filterMap protected theorem IsPrefix.reduceOption {l₁ l₂ : List (Option α)} (h : l₁ <+: l₂) : l₁.reduceOption <+: l₂.reduceOption := diff --git a/Mathlib/Data/List/Pairwise.lean b/Mathlib/Data/List/Pairwise.lean index be4f87bb84be6..be0725e086b7b 100644 --- a/Mathlib/Data/List/Pairwise.lean +++ b/Mathlib/Data/List/Pairwise.lean @@ -49,15 +49,6 @@ theorem Pairwise.forall (hR : Symmetric R) (hl : l.Pairwise R) : theorem Pairwise.set_pairwise (hl : Pairwise R l) (hr : Symmetric R) : { x | x ∈ l }.Pairwise R := hl.forall hr --- Porting note: Duplicate of `pairwise_map` but with `f` explicit. -@[deprecated "No deprecation message was provided." (since := "2024-02-25")] -theorem pairwise_map' (f : β → α) : - ∀ {l : List β}, Pairwise R (map f l) ↔ Pairwise (R on f) l - | [] => by simp only [map, Pairwise.nil] - | b :: l => by - simp only [map, pairwise_cons, mem_map, forall_exists_index, and_imp, - forall_apply_eq_imp_iff₂, pairwise_map] - theorem pairwise_of_reflexive_of_forall_ne {l : List α} {r : α → α → Prop} (hr : Reflexive r) (h : ∀ a ∈ l, ∀ b ∈ l, a ≠ b → r a b) : l.Pairwise r := by rw [pairwise_iff_forall_sublist] diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index aca81a62467f5..cf9436cf68118 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -184,11 +184,6 @@ theorem map_pure_sublist_sublists (l : List α) : map pure l <+ sublists l := by singleton_sublist.2 <| mem_map.2 ⟨[], mem_sublists.2 (nil_sublist _), by rfl⟩).trans ((append_sublist_append_right _).2 ih) -set_option linter.deprecated false in -@[deprecated map_pure_sublist_sublists (since := "2024-03-24")] -theorem map_ret_sublist_sublists (l : List α) : map List.ret l <+ sublists l := - map_pure_sublist_sublists l - /-! ### sublistsLen -/ /-- Auxiliary function to construct the list of all sublists of a given length. Given an diff --git a/Mathlib/Data/NNReal/Defs.lean b/Mathlib/Data/NNReal/Defs.lean index c773102bcabad..fb7ef70dca4ff 100644 --- a/Mathlib/Data/NNReal/Defs.lean +++ b/Mathlib/Data/NNReal/Defs.lean @@ -143,7 +143,6 @@ protected theorem coe_injective : Injective ((↑) : ℝ≥0 → ℝ) := Subtype @[simp, norm_cast] lemma coe_inj {r₁ r₂ : ℝ≥0} : (r₁ : ℝ) = r₂ ↔ r₁ = r₂ := NNReal.coe_injective.eq_iff -@[deprecated (since := "2024-02-03")] protected alias coe_eq := coe_inj @[simp, norm_cast] lemma coe_zero : ((0 : ℝ≥0) : ℝ) = 0 := rfl diff --git a/Mathlib/Data/Nat/Cast/Order/Basic.lean b/Mathlib/Data/Nat/Cast/Order/Basic.lean index eec03f2914acb..7d749878466c4 100644 --- a/Mathlib/Data/Nat/Cast/Order/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Order/Basic.lean @@ -33,9 +33,6 @@ theorem mono_cast : Monotone (Nat.cast : ℕ → α) := monotone_nat_of_le_succ fun n ↦ by rw [Nat.cast_succ]; exact le_add_of_nonneg_right zero_le_one -@[deprecated mono_cast (since := "2024-02-10")] -theorem cast_le_cast {a b : ℕ} (h : a ≤ b) : (a : α) ≤ b := mono_cast h - @[gcongr] theorem _root_.GCongr.natCast_le_natCast {a b : ℕ} (h : a ≤ b) : (a : α) ≤ b := mono_cast h diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index 5adc574882628..bde0bf0b51b40 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -740,7 +740,6 @@ lemma toDigitsCore_lens_eq_aux (b f : Nat) : specialize ih (n / b) (Nat.digitChar (n % b) :: l1) (Nat.digitChar (n % b) :: l2) simp only [List.length, congrArg (fun l ↦ l + 1) hlen] at ih exact ih trivial -@[deprecated (since := "2024-02-19")] alias to_digits_core_lens_eq_aux:= toDigitsCore_lens_eq_aux lemma toDigitsCore_lens_eq (b f : Nat) : ∀ (n : Nat) (c : Char) (tl : List Char), (Nat.toDigitsCore b f n (c :: tl)).length = (Nat.toDigitsCore b f n tl).length + 1 := by @@ -757,7 +756,6 @@ lemma toDigitsCore_lens_eq (b f : Nat) : ∀ (n : Nat) (c : Char) (tl : List Cha have lens_eq : (x :: (c :: tl)).length = (c :: x :: tl).length := by simp apply toDigitsCore_lens_eq_aux exact lens_eq -@[deprecated (since := "2024-02-19")] alias to_digits_core_lens_eq:= toDigitsCore_lens_eq lemma nat_repr_len_aux (n b e : Nat) (h_b_pos : 0 < b) : n < b ^ e.succ → n / b < b ^ e := by simp only [Nat.pow_succ] @@ -788,7 +786,6 @@ lemma toDigitsCore_length (b : Nat) (h : 2 <= b) (f n e : Nat) have _ : b ^ 1 = b := by simp only [Nat.pow_succ, pow_zero, Nat.one_mul] have _ : n < b := ‹b ^ 1 = b› ▸ hlt simp [(@Nat.div_eq_of_lt n b ‹n < b› : n / b = 0)] -@[deprecated (since := "2024-02-19")] alias to_digits_core_length := toDigitsCore_length /-- The core implementation of `Nat.repr` returns a String with length less than or equal to the number of digits in the decimal number (represented by `e`). For example, the decimal string diff --git a/Mathlib/Data/Rat/Cast/Defs.lean b/Mathlib/Data/Rat/Cast/Defs.lean index a5fc916543c00..5636ef6193774 100644 --- a/Mathlib/Data/Rat/Cast/Defs.lean +++ b/Mathlib/Data/Rat/Cast/Defs.lean @@ -114,8 +114,6 @@ theorem cast_intCast (n : ℤ) : ((n : ℚ) : α) = n := theorem cast_natCast (n : ℕ) : ((n : ℚ) : α) = n := by rw [← Int.cast_natCast, cast_intCast, Int.cast_natCast] -@[deprecated (since := "2024-03-21")] alias cast_coe_int := cast_intCast -@[deprecated (since := "2024-03-21")] alias cast_coe_nat := cast_natCast @[simp, norm_cast] lemma cast_ofNat (n : ℕ) [n.AtLeastTwo] : ((ofNat(n) : ℚ) : α) = (ofNat(n) : α) := by diff --git a/Mathlib/Data/Rat/Defs.lean b/Mathlib/Data/Rat/Defs.lean index cb77cd4efc0c3..f848e1c3e14a3 100644 --- a/Mathlib/Data/Rat/Defs.lean +++ b/Mathlib/Data/Rat/Defs.lean @@ -167,24 +167,14 @@ theorem lift_binop_eq (f : ℚ → ℚ → ℚ) (f₁ : ℤ → ℤ → ℤ → attribute [simp] divInt_add_divInt -@[deprecated divInt_add_divInt (since := "2024-03-18")] -theorem add_def'' {a b c d : ℤ} (b0 : b ≠ 0) (d0 : d ≠ 0) : - a /. b + c /. d = (a * d + c * b) /. (b * d) := divInt_add_divInt _ _ b0 d0 - attribute [simp] neg_divInt lemma neg_def (q : ℚ) : -q = -q.num /. q.den := by rw [← neg_divInt, num_divInt_den] @[simp] lemma divInt_neg (n d : ℤ) : n /. -d = -n /. d := divInt_neg' .. -@[deprecated (since := "2024-03-18")] alias divInt_neg_den := divInt_neg - attribute [simp] divInt_sub_divInt -@[deprecated divInt_sub_divInt (since := "2024-03-18")] -lemma sub_def'' {a b c d : ℤ} (b0 : b ≠ 0) (d0 : d ≠ 0) : - a /. b - c /. d = (a * d - c * b) /. (b * d) := divInt_sub_divInt _ _ b0 d0 - @[simp] lemma divInt_mul_divInt' (n₁ d₁ n₂ d₂ : ℤ) : (n₁ /. d₁) * (n₂ /. d₂) = (n₁ * n₂) /. (d₁ * d₂) := by obtain rfl | h₁ := eq_or_ne d₁ 0 @@ -270,17 +260,11 @@ protected theorem add_assoc : a + b + c = a + (b + c) := 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")] -lemma divInt_zero_one : 0 /. 1 = 0 := zero_divInt _ - @[simp] lemma divInt_one (n : ℤ) : n /. 1 = n := by simp [divInt, mkRat, normalize] @[simp] lemma mkRat_one (n : ℤ) : mkRat n 1 = n := by simp [mkRat_eq_divInt] lemma divInt_one_one : 1 /. 1 = 1 := by rw [divInt_one, intCast_one] -@[deprecated divInt_one (since := "2024-03-18")] -lemma divInt_neg_one_one : -1 /. 1 = -1 := by rw [divInt_one, intCast_neg, intCast_one] - protected theorem mul_assoc : a * b * c = a * (b * c) := numDenCasesOn' a fun n₁ d₁ h₁ => numDenCasesOn' b fun n₂ d₂ h₂ => diff --git a/Mathlib/Data/Real/Sqrt.lean b/Mathlib/Data/Real/Sqrt.lean index 5b1193b70b404..14d02f76715e2 100644 --- a/Mathlib/Data/Real/Sqrt.lean +++ b/Mathlib/Data/Real/Sqrt.lean @@ -62,11 +62,6 @@ lemma sqrt_le_iff_le_sq : sqrt x ≤ y ↔ x ≤ y ^ 2 := sqrt.to_galoisConnecti lemma le_sqrt_iff_sq_le : x ≤ sqrt y ↔ x ^ 2 ≤ y := (sqrt.symm.to_galoisConnection _ _).symm -@[deprecated (since := "2024-02-14")] alias sqrt_le_sqrt_iff := sqrt_le_sqrt -@[deprecated (since := "2024-02-14")] alias sqrt_lt_sqrt_iff := sqrt_lt_sqrt -@[deprecated (since := "2024-02-14")] alias sqrt_le_iff := sqrt_le_iff_le_sq -@[deprecated (since := "2024-02-14")] alias le_sqrt_iff := le_sqrt_iff_sq_le -@[deprecated (since := "2024-02-14")] alias sqrt_eq_iff_sq_eq := sqrt_eq_iff_eq_sq @[simp] lemma sqrt_eq_zero : sqrt x = 0 ↔ x = 0 := by simp [sqrt_eq_iff_eq_sq] diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 64f61f2b25e36..0a7d5df2df1b0 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -514,7 +514,6 @@ theorem subset_eq_empty {s t : Set α} (h : t ⊆ s) (e : s = ∅) : t = ∅ := theorem forall_mem_empty {p : α → Prop} : (∀ x ∈ (∅ : Set α), p x) ↔ True := iff_true_intro fun _ => False.elim -@[deprecated (since := "2024-03-23")] alias ball_empty_iff := forall_mem_empty instance (α : Type u) : IsEmpty.{u + 1} (↥(∅ : Set α)) := ⟨fun x => x.2⟩ @@ -843,10 +842,6 @@ theorem union_inter_distrib_left (s t u : Set α) : s ∪ t ∩ u = (s ∪ t) theorem inter_union_distrib_right (s t u : Set α) : s ∩ t ∪ u = (s ∪ u) ∩ (t ∪ u) := sup_inf_right _ _ _ -@[deprecated (since := "2024-03-22")] alias inter_distrib_left := inter_union_distrib_left -@[deprecated (since := "2024-03-22")] alias inter_distrib_right := union_inter_distrib_right -@[deprecated (since := "2024-03-22")] alias union_distrib_left := union_inter_distrib_left -@[deprecated (since := "2024-03-22")] alias union_distrib_right := inter_union_distrib_right theorem union_union_distrib_left (s t u : Set α) : s ∪ (t ∪ u) = s ∪ t ∪ (s ∪ u) := sup_sup_distrib_left _ _ _ @@ -979,12 +974,10 @@ theorem forall_insert_of_forall {P : α → Prop} {a : α} {s : Set α} (H : ∀ theorem exists_mem_insert {P : α → Prop} {a : α} {s : Set α} : (∃ x ∈ insert a s, P x) ↔ (P a ∨ ∃ x ∈ s, P x) := by simp [mem_insert_iff, or_and_right, exists_and_left, exists_or] -@[deprecated (since := "2024-03-23")] alias bex_insert_iff := exists_mem_insert theorem forall_mem_insert {P : α → Prop} {a : α} {s : Set α} : (∀ x ∈ insert a s, P x) ↔ P a ∧ ∀ x ∈ s, P x := forall₂_or_left.trans <| and_congr_left' forall_eq -@[deprecated (since := "2024-03-23")] alias ball_insert_iff := forall_mem_insert /-- Inserting an element to a set is equivalent to the option type. -/ def subtypeInsertEquivOption diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index a98696d44ae14..a36c4d6171444 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -186,11 +186,6 @@ variable {f : α → β} {s t : Set α} -- Porting note: `Set.image` is already defined in `Data.Set.Defs` -@[deprecated mem_image (since := "2024-03-23")] -theorem mem_image_iff_bex {f : α → β} {s : Set α} {y : β} : - y ∈ f '' s ↔ ∃ (x : _) (_ : x ∈ s), f x = y := - bex_def.symm - theorem image_eta (f : α → β) : f '' s = (fun x => f x) '' s := rfl @@ -208,18 +203,6 @@ theorem forall_mem_image {f : α → β} {s : Set α} {p : β → Prop} : theorem exists_mem_image {f : α → β} {s : Set α} {p : β → Prop} : (∃ y ∈ f '' s, p y) ↔ ∃ x ∈ s, p (f x) := by simp -@[deprecated (since := "2024-02-21")] alias ball_image_iff := forall_mem_image -@[deprecated (since := "2024-02-21")] alias bex_image_iff := exists_mem_image -@[deprecated (since := "2024-02-21")] alias ⟨_, ball_image_of_ball⟩ := forall_mem_image - -@[deprecated forall_mem_image (since := "2024-02-21")] -theorem mem_image_elim {f : α → β} {s : Set α} {C : β → Prop} (h : ∀ x : α, x ∈ s → C (f x)) : - ∀ {y : β}, y ∈ f '' s → C y := forall_mem_image.2 h _ - -@[deprecated forall_mem_image (since := "2024-02-21")] -theorem mem_image_elim_on {f : α → β} {s : Set α} {C : β → Prop} {y : β} (h_y : y ∈ f '' s) - (h : ∀ x : α, x ∈ s → C (f x)) : C y := forall_mem_image.2 h _ h_y - -- Porting note: used to be `safe` @[congr] theorem image_congr {f g : α → β} {s : Set α} (h : ∀ a ∈ s, f a = g a) : f '' s = g '' s := by @@ -574,7 +557,6 @@ variable {f : ι → α} {s t : Set α} theorem forall_mem_range {p : α → Prop} : (∀ a ∈ range f, p a) ↔ ∀ i, p (f i) := by simp -@[deprecated (since := "2024-02-21")] alias forall_range_iff := forall_mem_range theorem forall_subtype_range_iff {p : range f → Prop} : (∀ a : range f, p a) ↔ ∀ i, p ⟨f i, mem_range_self _⟩ := @@ -584,9 +566,6 @@ theorem forall_subtype_range_iff {p : range f → Prop} : theorem exists_range_iff {p : α → Prop} : (∃ a ∈ range f, p a) ↔ ∃ i, p (f i) := by simp -@[deprecated (since := "2024-03-10")] -alias exists_range_iff' := exists_range_iff - theorem exists_subtype_range_iff {p : range f → Prop} : (∃ a : range f, p a) ↔ ∃ i, p ⟨f i, mem_range_self _⟩ := ⟨fun ⟨⟨a, i, hi⟩, ha⟩ => by diff --git a/Mathlib/Data/ZMod/Quotient.lean b/Mathlib/Data/ZMod/Quotient.lean index bcbc87a1d2aa0..1f9876b6be3b6 100644 --- a/Mathlib/Data/ZMod/Quotient.lean +++ b/Mathlib/Data/ZMod/Quotient.lean @@ -142,9 +142,6 @@ theorem orbitZPowersEquiv_symm_apply (k : ZMod (minimalPeriod (a • ·) b)) : (⟨a, mem_zpowers a⟩ : zpowers a) ^ (cast k : ℤ) • ⟨b, mem_orbit_self b⟩ := rfl -@[deprecated (since := "2024-02-21")] -alias _root_.AddAction.orbit_zmultiples_equiv_symm_apply := orbitZMultiplesEquiv_symm_apply - theorem orbitZPowersEquiv_symm_apply' (k : ℤ) : (orbitZPowersEquiv a b).symm k = (⟨a, mem_zpowers a⟩ : zpowers a) ^ k • ⟨b, mem_orbit_self b⟩ := by diff --git a/Mathlib/Deprecated/Cardinal/PartENat.lean b/Mathlib/Deprecated/Cardinal/PartENat.lean index e2c263a8cfb2e..47e21caf7cb31 100644 --- a/Mathlib/Deprecated/Cardinal/PartENat.lean +++ b/Mathlib/Deprecated/Cardinal/PartENat.lean @@ -48,9 +48,6 @@ theorem toPartENat_eq_top {c : Cardinal} : theorem toPartENat_apply_of_aleph0_le {c : Cardinal} (h : ℵ₀ ≤ c) : toPartENat c = ⊤ := congr_arg PartENat.ofENat (toENat_eq_top.2 h) -@[deprecated (since := "2024-02-15")] -alias toPartENat_cast := toPartENat_natCast - @[simp] theorem mk_toPartENat_of_infinite [h : Infinite α] : toPartENat #α = ⊤ := toPartENat_apply_of_aleph0_le (infinite_iff.1 h) @@ -62,7 +59,6 @@ theorem aleph0_toPartENat : toPartENat ℵ₀ = ⊤ := theorem toPartENat_surjective : Surjective toPartENat := fun x => PartENat.casesOn x ⟨ℵ₀, toPartENat_apply_of_aleph0_le le_rfl⟩ fun n => ⟨n, toPartENat_natCast n⟩ -@[deprecated (since := "2024-02-15")] alias toPartENat_eq_top_iff_le_aleph0 := toPartENat_eq_top theorem toPartENat_strictMonoOn : StrictMonoOn toPartENat (Set.Iic ℵ₀) := PartENat.withTopOrderIso.symm.strictMono.comp_strictMonoOn toENat_strictMonoOn diff --git a/Mathlib/Deprecated/Submonoid.lean b/Mathlib/Deprecated/Submonoid.lean index 2e9618d300fb9..104e7376fd117 100644 --- a/Mathlib/Deprecated/Submonoid.lean +++ b/Mathlib/Deprecated/Submonoid.lean @@ -184,7 +184,6 @@ theorem IsSubmonoid.pow_mem {a : M} (hs : IsSubmonoid s) (h : a ∈ s) : ∀ {n the `AddSubmonoid`."] theorem IsSubmonoid.powers_subset {a : M} (hs : IsSubmonoid s) (h : a ∈ s) : powers a ⊆ s := fun _ ⟨_, hx⟩ => hx ▸ hs.pow_mem h -@[deprecated (since := "2024-02-21")] alias IsSubmonoid.power_subset := IsSubmonoid.powers_subset end powers diff --git a/Mathlib/FieldTheory/PrimitiveElement.lean b/Mathlib/FieldTheory/PrimitiveElement.lean index 6b12f4601768f..b8174c3c41d64 100644 --- a/Mathlib/FieldTheory/PrimitiveElement.lean +++ b/Mathlib/FieldTheory/PrimitiveElement.lean @@ -286,9 +286,6 @@ theorem FiniteDimensional.of_finite_intermediateField rw [htop] at hfin exact topEquiv.toLinearEquiv.finiteDimensional -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_finite_intermediateField := FiniteDimensional.of_finite_intermediateField - theorem exists_primitive_element_of_finite_intermediateField [Finite (IntermediateField F E)] (K : IntermediateField F E) : ∃ α : E, F⟮α⟯ = K := by haveI := FiniteDimensional.of_finite_intermediateField F E @@ -307,9 +304,6 @@ theorem FiniteDimensional.of_exists_primitive_element [Algebra.IsAlgebraic F E] rw [hprim] at hfin exact topEquiv.toLinearEquiv.finiteDimensional -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_exists_primitive_element := FiniteDimensional.of_exists_primitive_element - -- A finite simple extension has only finitely many intermediate fields theorem finite_intermediateField_of_exists_primitive_element [Algebra.IsAlgebraic F E] (h : ∃ α : E, F⟮α⟯ = ⊤) : Finite (IntermediateField F E) := by diff --git a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean index 782d2803dcdb5..1531fd8dcb8b4 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean @@ -136,9 +136,6 @@ nonrec theorem angle_le_pi (p1 p2 p3 : P) : ∠ p1 p2 p3 ≤ π := /-- 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 -@[deprecated (since := "2024-02-14")] alias angle_eq_left := angle_self_left -@[deprecated (since := "2024-02-14")] alias angle_eq_right := angle_self_right -@[deprecated (since := "2024-02-14")] alias angle_eq_of_ne := angle_self_of_ne /-- If the angle ∠ABC at a point is π, the angle ∠BAC is 0. -/ theorem angle_eq_zero_of_angle_eq_pi_left {p1 p2 p3 : P} (h : ∠ p1 p2 p3 = π) : ∠ p2 p1 p3 = 0 := by diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index d3b381f0c0511..4bf23adb82cdd 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -303,7 +303,6 @@ alias closedEmbedding := isClosedEmbedding theorem isClosed_range : IsClosed (range I) := I.isClosedEmbedding.isClosed_range -@[deprecated (since := "2024-03-17")] alias closed_range := isClosed_range theorem range_eq_closure_interior : range I = closure (interior (range I)) := Subset.antisymm I.range_subset_closure_interior I.isClosed_range.closure_interior_subset diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index a5f1e2aead8f6..514fd7841f060 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -263,9 +263,6 @@ protected lemma IsOfFinOrder.mem_powers_iff_mem_range_orderOf [DecidableEq G] protected lemma IsOfFinOrder.powers_eq_image_range_orderOf [DecidableEq G] (hx : IsOfFinOrder x) : (Submonoid.powers x : Set G) = (Finset.range (orderOf x)).image (x ^ ·) := Set.ext fun _ ↦ hx.mem_powers_iff_mem_range_orderOf -@[deprecated (since := "2024-02-21")] -alias IsOfFinAddOrder.powers_eq_image_range_orderOf := - IsOfFinAddOrder.multiples_eq_image_range_addOrderOf @[to_additive] theorem pow_eq_one_iff_modEq : x ^ n = 1 ↔ n ≡ 0 [MOD orderOf x] := by @@ -1089,19 +1086,14 @@ variable [Monoid α] [Monoid β] {x : α × β} {a : α} {b : β} @[to_additive] protected theorem Prod.orderOf (x : α × β) : orderOf x = (orderOf x.1).lcm (orderOf x.2) := minimalPeriod_prod_map _ _ _ -@[deprecated (since := "2024-02-21")] alias Prod.add_orderOf := Prod.addOrderOf @[to_additive] theorem orderOf_fst_dvd_orderOf : orderOf x.1 ∣ orderOf x := minimalPeriod_fst_dvd -@[deprecated (since := "2024-02-21")] -alias add_orderOf_fst_dvd_add_orderOf := addOrderOf_fst_dvd_addOrderOf @[to_additive] theorem orderOf_snd_dvd_orderOf : orderOf x.2 ∣ orderOf x := minimalPeriod_snd_dvd -@[deprecated (since := "2024-02-21")] alias -add_orderOf_snd_dvd_add_orderOf := addOrderOf_snd_dvd_addOrderOf @[to_additive] theorem IsOfFinOrder.fst {x : α × β} (hx : IsOfFinOrder x) : IsOfFinOrder x.1 := diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index 5ea4adad723cf..b469bc7d71d9f 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -105,8 +105,6 @@ theorem MonoidHom.map_cyclic [h : IsCyclic G] (σ : G →* G) : refine ⟨m, fun g => ?_⟩ obtain ⟨n, rfl⟩ := hG g rw [MonoidHom.map_zpow, ← hm, ← zpow_mul, ← zpow_mul'] -@[deprecated (since := "2024-02-21")] alias -MonoidAddHom.map_add_cyclic := AddMonoidHom.map_addCyclic @[to_additive] lemma isCyclic_iff_exists_orderOf_eq_natCard [Finite α] : @@ -132,9 +130,6 @@ theorem isCyclic_of_orderOf_eq_card [Finite α] (x : α) (hx : orderOf x = Nat.c IsCyclic α := isCyclic_iff_exists_orderOf_eq_natCard.mpr ⟨x, hx⟩ -@[deprecated (since := "2024-02-21")] -alias isAddCyclic_of_orderOf_eq_card := isAddCyclic_of_addOrderOf_eq_card - @[to_additive] theorem Subgroup.eq_bot_or_eq_top_of_prime_card (H : Subgroup G) [hp : Fact (Nat.card G).Prime] : H = ⊥ ∨ H = ⊤ := by @@ -324,8 +319,6 @@ theorem IsCyclic.card_pow_eq_one_le [DecidableEq α] [Fintype α] [IsCyclic α] rw [Nat.mul_div_cancel_left _ (gcd_pos_of_pos_left _ hn0), gcd_mul_left_left, hm, Nat.mul_div_cancel _ hm0] exact le_of_dvd hn0 (Nat.gcd_dvd_left _ _) -@[deprecated (since := "2024-02-21")] -alias IsAddCyclic.card_pow_eq_one_le := IsAddCyclic.card_nsmul_eq_zero_le end Classical @@ -478,9 +471,6 @@ theorem isCyclic_of_card_pow_eq_one_le : IsCyclic α := let ⟨x, hx⟩ := this isCyclic_of_orderOf_eq_card x (Finset.mem_filter.1 hx).2 -@[deprecated (since := "2024-02-21")] -alias isAddCyclic_of_card_pow_eq_one_le := isAddCyclic_of_card_nsmul_eq_zero_le - end Totient @[to_additive] @@ -488,9 +478,6 @@ lemma IsCyclic.card_orderOf_eq_totient [IsCyclic α] [Fintype α] {d : ℕ} (hd #{a : α | orderOf a = d} = totient d := by classical apply card_orderOf_eq_totient_aux₂ (fun n => IsCyclic.card_pow_eq_one_le) hd -@[deprecated (since := "2024-02-21")] -alias IsAddCyclic.card_orderOf_eq_totient := IsAddCyclic.card_addOrderOf_eq_totient - /-- A finite group of prime order is simple. -/ @[to_additive "A finite group of prime order is simple."] theorem isSimpleGroup_of_prime_card {p : ℕ} [hp : Fact p.Prime] @@ -532,9 +519,6 @@ theorem commutative_of_cyclic_center_quotient [IsCyclic G'] (f : G →* G') (hf _ = y ^ m * y ^ n * y ^ (-m) * (y ^ (-n) * b * a) := by rw [mem_center_iff.1 hb] _ = b * a := by group -@[deprecated (since := "2024-02-21")] -alias commutative_of_add_cyclic_center_quotient := commutative_of_addCyclic_center_quotient - /-- A group is commutative if the quotient by the center is cyclic. -/ @[to_additive "A group is commutative if the quotient by the center is cyclic."] diff --git a/Mathlib/LinearAlgebra/FiniteDimensional.lean b/Mathlib/LinearAlgebra/FiniteDimensional.lean index 3795257b53d17..1a9987e8e47be 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional.lean @@ -227,10 +227,6 @@ theorem LinearIndependent.span_eq_top_of_card_eq_finrank {ι : Type*} [Nonempty have : FiniteDimensional K V := .of_finrank_pos <| card_eq ▸ Fintype.card_pos lin_ind.span_eq_top_of_card_eq_finrank' card_eq -@[deprecated (since := "2024-02-14")] -alias span_eq_top_of_linearIndependent_of_card_eq_finrank := - LinearIndependent.span_eq_top_of_card_eq_finrank - /-- A linear independent family of `finrank K V` vectors forms a basis. -/ @[simps! repr_apply] noncomputable def basisOfLinearIndependentOfCardEqFinrank {ι : Type*} [Nonempty ι] [Fintype ι] diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index 16ac3a373d3e2..66b253e50a6c7 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -322,21 +322,12 @@ theorem FiniteDimensional.of_rank_eq_nat {n : ℕ} (h : Module.rank K V = n) : FiniteDimensional K V := Module.finite_of_rank_eq_nat h -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_rank_eq_nat := FiniteDimensional.of_rank_eq_nat - theorem FiniteDimensional.of_rank_eq_zero (h : Module.rank K V = 0) : FiniteDimensional K V := Module.finite_of_rank_eq_zero h -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_rank_eq_zero := FiniteDimensional.of_rank_eq_zero - theorem FiniteDimensional.of_rank_eq_one (h : Module.rank K V = 1) : FiniteDimensional K V := Module.finite_of_rank_eq_one h -@[deprecated (since := "2024-02-02")] -alias finiteDimensional_of_rank_eq_one := FiniteDimensional.of_rank_eq_one - variable (K V) instance finiteDimensional_bot : FiniteDimensional K (⊥ : Submodule K V) := diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 409b30a71b81e..56e15a7a87dc5 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -375,8 +375,6 @@ theorem forall_mem_comm {α β} [Membership α β] {s : β} {p : α → α → P (∀ a (_ : a ∈ s) b (_ : b ∈ s), p a b) ↔ ∀ a b, a ∈ s → b ∈ s → p a b := forall_cond_comm -@[deprecated (since := "2024-03-23")] alias ball_cond_comm := forall_cond_comm -@[deprecated (since := "2024-03-23")] alias ball_mem_comm := forall_mem_comm lemma ne_of_eq_of_ne {α : Sort*} {a b c : α} (h₁ : a = b) (h₂ : b ≠ c) : a ≠ c := h₁.symm ▸ h₂ lemma ne_of_ne_of_eq {α : Sort*} {a b c : α} (h₁ : a ≠ b) (h₂ : b = c) : a ≠ c := h₂ ▸ h₁ @@ -783,21 +781,12 @@ theorem BAll.imp_left (H : ∀ x, p x → q x) (h₁ : ∀ x, q x → r x) (x) ( theorem BEx.imp_left (H : ∀ x, p x → q x) : (∃ (x : _) (_ : p x), r x) → ∃ (x : _) (_ : q x), r x | ⟨x, hp, hr⟩ => ⟨x, H _ hp, hr⟩ -@[deprecated id (since := "2024-03-23")] -theorem ball_of_forall (h : ∀ x, p x) (x) : p x := h x - -@[deprecated forall_imp (since := "2024-03-23")] -theorem forall_of_ball (H : ∀ x, p x) (h : ∀ x, p x → q x) (x) : q x := h x <| H x - theorem exists_mem_of_exists (H : ∀ x, p x) : (∃ x, q x) → ∃ (x : _) (_ : p x), q x | ⟨x, hq⟩ => ⟨x, H x, hq⟩ theorem exists_of_exists_mem : (∃ (x : _) (_ : p x), q x) → ∃ x, q x | ⟨x, _, hq⟩ => ⟨x, hq⟩ -@[deprecated (since := "2024-03-23")] alias bex_of_exists := exists_mem_of_exists -@[deprecated (since := "2024-03-23")] alias exists_of_bex := exists_of_exists_mem -@[deprecated (since := "2024-03-23")] alias bex_imp := exists₂_imp theorem not_exists_mem : (¬∃ x h, P x h) ↔ ∀ x h, ¬P x h := exists₂_imp diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean index adcc5cb835bb0..7af987e92ef57 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean @@ -195,9 +195,6 @@ theorem aemeasurable_coe_nnreal_real_iff {f : α → ℝ≥0} {μ : Measure α} AEMeasurable (fun x => f x : α → ℝ) μ ↔ AEMeasurable f μ := ⟨fun h ↦ by simpa only [Real.toNNReal_coe] using h.real_toNNReal, AEMeasurable.coe_nnreal_real⟩ -@[deprecated (since := "2024-03-02")] -alias aEMeasurable_coe_nnreal_real_iff := aemeasurable_coe_nnreal_real_iff - /-- The set of finite `ℝ≥0∞` numbers is `MeasurableEquiv` to `ℝ≥0`. -/ def MeasurableEquiv.ennrealEquivNNReal : { r : ℝ≥0∞ | r ≠ ∞ } ≃ᵐ ℝ≥0 := ENNReal.neTopHomeomorphNNReal.toMeasurableEquiv @@ -283,17 +280,11 @@ theorem measurable_of_tendsto' {ι : Type*} {f : ι → α → ℝ≥0∞} {g : show Measurable fun y => liminf (fun n => (f (x n) y : ℝ≥0∞)) atTop exact .liminf fun n => hf (x n) -@[deprecated (since := "2024-03-09")] alias -_root_.measurable_of_tendsto_ennreal' := ENNReal.measurable_of_tendsto' - /-- A sequential limit of measurable `ℝ≥0∞` valued functions is measurable. -/ theorem measurable_of_tendsto {f : ℕ → α → ℝ≥0∞} {g : α → ℝ≥0∞} (hf : ∀ i, Measurable (f i)) (lim : Tendsto f atTop (𝓝 g)) : Measurable g := measurable_of_tendsto' atTop hf lim -@[deprecated (since := "2024-03-09")] alias -_root_.measurable_of_tendsto_ennreal := ENNReal.measurable_of_tendsto - /-- A limit (over a general filter) of a.e.-measurable `ℝ≥0∞` valued functions is a.e.-measurable. -/ lemma aemeasurable_of_tendsto' {ι : Type*} {f : ι → α → ℝ≥0∞} {g : α → ℝ≥0∞} @@ -453,17 +444,11 @@ theorem measurable_of_tendsto' {ι} {f : ι → α → ℝ≥0} {g : α → ℝ rw [tendsto_pi_nhds] at lim ⊢ exact fun x => (ENNReal.continuous_coe.tendsto (g x)).comp (lim x) -@[deprecated (since := "2024-03-09")] alias -_root_.measurable_of_tendsto_nnreal' := NNReal.measurable_of_tendsto' - /-- A sequential limit of measurable `ℝ≥0` valued functions is measurable. -/ theorem measurable_of_tendsto {f : ℕ → α → ℝ≥0} {g : α → ℝ≥0} (hf : ∀ i, Measurable (f i)) (lim : Tendsto f atTop (𝓝 g)) : Measurable g := measurable_of_tendsto' atTop hf lim -@[deprecated (since := "2024-03-09")] alias -_root_.measurable_of_tendsto_nnreal := NNReal.measurable_of_tendsto - end NNReal namespace EReal diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space.lean index 30d728e7e06fc..0005ef1d22c72 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space.lean @@ -173,9 +173,6 @@ theorem HasFiniteIntegral.of_finite [Finite α] [IsFiniteMeasure μ] {f : α → let ⟨_⟩ := nonempty_fintype α hasFiniteIntegral_of_bounded <| ae_of_all μ <| norm_le_pi_norm f -@[deprecated (since := "2024-02-05")] -alias hasFiniteIntegral_of_fintype := HasFiniteIntegral.of_finite - theorem HasFiniteIntegral.mono_measure {f : α → β} (h : HasFiniteIntegral f ν) (hμ : μ ≤ ν) : HasFiniteIntegral f μ := lt_of_le_of_lt (lintegral_mono' hμ le_rfl) h @@ -474,7 +471,6 @@ lemma Integrable.of_finite [Finite α] [MeasurableSingletonClass α] [IsFiniteMe @[deprecated Integrable.of_finite (since := "2024-10-05"), nolint deprecatedNoSince] lemma Integrable.of_isEmpty [IsEmpty α] {f : α → β} : Integrable f μ := .of_finite -@[deprecated (since := "2024-02-05")] alias integrable_of_fintype := Integrable.of_finite theorem Memℒp.integrable_norm_rpow {f : α → β} {p : ℝ≥0∞} (hf : Memℒp f p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : Integrable (fun x : α => ‖f x‖ ^ p.toReal) μ := by diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 05df7e2bf5776..3d030c368a939 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -77,7 +77,6 @@ def ofFinite [Finite α] [MeasurableSingletonClass α] (f : α → β) : α → measurableSet_fiber' x := (toFinite (f ⁻¹' {x})).measurableSet finite_range' := Set.finite_range f -@[deprecated (since := "2024-02-05")] alias ofFintype := ofFinite /-- Simple function defined on the empty type. -/ def ofIsEmpty [IsEmpty α] : α →ₛ β := ofFinite isEmptyElim diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index 757c23abb972f..3fcd3f10b57d4 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -130,14 +130,6 @@ theorem StronglyMeasurable.of_finite [Finite α] {_ : MeasurableSpace α} {f : α → β} : StronglyMeasurable f := ⟨fun _ => SimpleFunc.ofFinite f, fun _ => tendsto_const_nhds⟩ -@[deprecated (since := "2024-02-05")] -alias stronglyMeasurable_of_fintype := StronglyMeasurable.of_finite - -@[deprecated StronglyMeasurable.of_finite (since := "2024-02-06")] -theorem stronglyMeasurable_of_isEmpty [IsEmpty α] {_ : MeasurableSpace α} [TopologicalSpace β] - (f : α → β) : StronglyMeasurable f := - .of_finite - theorem stronglyMeasurable_const {α β} {_ : MeasurableSpace α} [TopologicalSpace β] {b : β} : StronglyMeasurable fun _ : α => b := ⟨fun _ => SimpleFunc.const α b, fun _ => tendsto_const_nhds⟩ @@ -1162,7 +1154,6 @@ theorem mono_measure {ν : Measure α} (hf : AEStronglyMeasurable f μ) (h : ν protected lemma mono_ac (h : ν ≪ μ) (hμ : AEStronglyMeasurable f μ) : AEStronglyMeasurable f ν := let ⟨g, hg, hg'⟩ := hμ; ⟨g, hg, h.ae_eq hg'⟩ -@[deprecated (since := "2024-02-15")] protected alias mono' := AEStronglyMeasurable.mono_ac theorem mono_set {s t} (h : s ⊆ t) (ht : AEStronglyMeasurable f (μ.restrict t)) : AEStronglyMeasurable f (μ.restrict s) := diff --git a/Mathlib/MeasureTheory/Integral/PeakFunction.lean b/Mathlib/MeasureTheory/Integral/PeakFunction.lean index 20b7988c88759..ea9c6d69cf046 100644 --- a/Mathlib/MeasureTheory/Integral/PeakFunction.lean +++ b/Mathlib/MeasureTheory/Integral/PeakFunction.lean @@ -83,10 +83,6 @@ theorem integrableOn_peak_smul_of_integrableOn_of_tendsto convert A.union B simp only [diff_union_inter] -@[deprecated (since := "2024-02-20")] -alias integrableOn_peak_smul_of_integrableOn_of_continuousWithinAt := - integrableOn_peak_smul_of_integrableOn_of_tendsto - /-- 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`. @@ -176,10 +172,6 @@ theorem tendsto_setIntegral_peak_smul_of_integrableOn_of_tendsto_aux _ ≤ ‖∫ x in s \ u, φ i x • g x ∂μ‖ + ‖∫ x in s ∩ u, φ i x • g x ∂μ‖ := norm_add_le _ _ _ ≤ (δ * ∫ x in s, ‖g x‖ ∂μ) + 2 * δ := add_le_add C B -@[deprecated (since := "2024-02-20")] -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 @@ -220,10 +212,6 @@ theorem tendsto_setIntegral_peak_smul_of_integrableOn_of_tendsto rw [restrict_restrict ht, inter_eq_left.mpr hts] exact .of_integral_ne_zero (fun h ↦ by simp [h] at h'i) -@[deprecated (since := "2024-02-20")] -alias tendsto_setIntegral_peak_smul_of_integrableOn_of_continuousWithinAt := - tendsto_setIntegral_peak_smul_of_integrableOn_of_tendsto - /-- 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`. -/ diff --git a/Mathlib/MeasureTheory/Integral/SetToL1.lean b/Mathlib/MeasureTheory/Integral/SetToL1.lean index 7c6e7f8e1b323..9c3a4587da7c6 100644 --- a/Mathlib/MeasureTheory/Integral/SetToL1.lean +++ b/Mathlib/MeasureTheory/Integral/SetToL1.lean @@ -527,9 +527,6 @@ theorem norm_setToSimpleFunc_le_sum_opNorm {m : MeasurableSpace α} (T : Set α _ ≤ ∑ x ∈ f.range, ‖T (f ⁻¹' {x})‖ * ‖x‖ := by refine Finset.sum_le_sum fun b _ => ?_; simp_rw [ContinuousLinearMap.le_opNorm] -@[deprecated (since := "2024-02-02")] -alias norm_setToSimpleFunc_le_sum_op_norm := norm_setToSimpleFunc_le_sum_opNorm - theorem norm_setToSimpleFunc_le_sum_mul_norm (T : Set α → F →L[ℝ] F') {C : ℝ} (hT_norm : ∀ s, MeasurableSet s → ‖T s‖ ≤ C * (μ s).toReal) (f : α →ₛ F) : ‖f.setToSimpleFunc T‖ ≤ C * ∑ x ∈ f.range, (μ (f ⁻¹' {x})).toReal * ‖x‖ := diff --git a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean index e9d2e87202f9e..a4ed925b7ac08 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean @@ -390,12 +390,6 @@ lemma haarScalarFactor_pos_of_isHaarMeasure (μ' μ : Measure G) [IsHaarMeasure [IsHaarMeasure μ'] : 0 < haarScalarFactor μ' μ := pos_iff_ne_zero.2 (fun H ↦ by simpa [H] using haarScalarFactor_eq_mul μ' μ μ') -@[deprecated (since := "2024-02-12")] -alias haarScalarFactor_pos_of_isOpenPosMeasure := haarScalarFactor_pos_of_isHaarMeasure - -@[deprecated (since := "2024-02-12")] -alias addHaarScalarFactor_pos_of_isOpenPosMeasure := addHaarScalarFactor_pos_of_isAddHaarMeasure - /-! ### Uniqueness of measure of sets with compact closure @@ -681,13 +675,6 @@ lemma isHaarMeasure_eq_of_isProbabilityMeasure [LocallyCompactSpace G] (μ' μ : ext s _hs simp [A s, ← Z] -@[deprecated (since := "2024-02-12")] -alias haarScalarFactor_eq_one_of_isProbabilityMeasure := isHaarMeasure_eq_of_isProbabilityMeasure - -@[deprecated (since := "2024-02-12")] -alias addHaarScalarFactor_eq_one_of_isProbabilityMeasure := - isAddHaarMeasure_eq_of_isProbabilityMeasure - /-! ### Uniqueness of measure of open sets @@ -896,8 +883,6 @@ lemma isMulLeftInvariant_eq_smul [LocallyCompactSpace G] [SecondCountableTopolog -- one could use as well `isMulLeftInvariant_eq_smul_of_innerRegular`, as in a -- second countable topological space all Haar measures are regular and inner regular -@[deprecated (since := "2024-02-12")] alias isHaarMeasure_eq_smul := isMulLeftInvariant_eq_smul -@[deprecated (since := "2024-02-12")] alias isAddHaarMeasure_eq_smul := isAddLeftInvariant_eq_smul /-- An invariant σ-finite measure is absolutely continuous with respect to a Haar measure in a second countable group. -/ diff --git a/Mathlib/NumberTheory/LSeries/Dirichlet.lean b/Mathlib/NumberTheory/LSeries/Dirichlet.lean index 32e114a35c3d6..96724285e05ad 100644 --- a/Mathlib/NumberTheory/LSeries/Dirichlet.lean +++ b/Mathlib/NumberTheory/LSeries/Dirichlet.lean @@ -264,9 +264,6 @@ lemma LSeries_zeta_eq : L ↗ζ = L 1 := by theorem LSeriesSummable_zeta_iff {s : ℂ} : LSeriesSummable (ζ ·) s ↔ 1 < s.re := (LSeriesSummable_congr s const_one_eq_zeta).symm.trans <| LSeriesSummable_one_iff -@[deprecated (since := "2024-03-29")] -alias zeta_LSeriesSummable_iff_one_lt_re := LSeriesSummable_zeta_iff - /-- The abscissa of (absolute) convergence of the arithmetic function `ζ` is `1`. -/ lemma abscissaOfAbsConv_zeta : abscissaOfAbsConv ↗ζ = 1 := by rw [abscissaOfAbsConv_congr (g := 1) fun hn ↦ by simp [hn], abscissaOfAbsConv_one] diff --git a/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean b/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean index 9e0eff961d50c..fc253159650ba 100644 --- a/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean +++ b/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean @@ -34,7 +34,6 @@ theorem IsGδ.setOf_liouville : IsGδ { x | Liouville x } := by refine isOpen_iUnion fun a => isOpen_iUnion fun b => isOpen_iUnion fun _hb => ?_ exact isOpen_ball.inter isClosed_singleton.isOpen_compl -@[deprecated (since := "2024-02-15")] alias isGδ_setOf_liouville := IsGδ.setOf_liouville theorem setOf_liouville_eq_irrational_inter_iInter_iUnion : { x | Liouville x } = diff --git a/Mathlib/RingTheory/Flat/Basic.lean b/Mathlib/RingTheory/Flat/Basic.lean index 66b8a73a95d3f..0923edaa450fb 100644 --- a/Mathlib/RingTheory/Flat/Basic.lean +++ b/Mathlib/RingTheory/Flat/Basic.lean @@ -107,9 +107,6 @@ theorem iff_rTensor_injective' : rewrite [← rTensor_comp_apply] at hx₀ rw [(injective_iff_map_eq_zero _).mp (h hfg) y hx₀, LinearMap.map_zero] -@[deprecated (since := "2024-03-29")] -alias lTensor_inj_iff_rTensor_inj := LinearMap.lTensor_inj_iff_rTensor_inj - /-- The `lTensor`-variant of `iff_rTensor_injective`. . -/ theorem iff_lTensor_injective : Module.Flat R M ↔ ∀ ⦃I : Ideal R⦄ (_ : I.FG), Function.Injective (lTensor M I.subtype) := by @@ -259,9 +256,6 @@ theorem rTensor_preserves_injective_linearMap {N' : Type*} [AddCommGroup N'] [Mo Function.Injective (L.rTensor M) := rTensor_injective_iff_lcomp_surjective.2 ((iff_characterModule_baer.1 h).extension_property _ hL) -@[deprecated (since := "2024-03-29")] -alias preserves_injective_linearMap := rTensor_preserves_injective_linearMap - instance {S} [CommRing S] [Algebra R S] [Module S M] [IsScalarTower R S M] [Flat S M] [Flat R N] : Flat S (M ⊗[R] N) := (iff_rTensor_injective' _ _).mpr fun I ↦ by diff --git a/Mathlib/RingTheory/Ideal/Quotient/Operations.lean b/Mathlib/RingTheory/Ideal/Quotient/Operations.lean index 488ffed5743b2..f18039405d13b 100644 --- a/Mathlib/RingTheory/Ideal/Quotient/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Quotient/Operations.lean @@ -445,11 +445,6 @@ def quotientKerAlgEquivOfRightInverse {f : A →ₐ[R₁] B} {g : B → A} { RingHom.quotientKerEquivOfRightInverse hf, kerLiftAlg f with } -@[deprecated (since := "2024-02-27")] -alias quotientKerAlgEquivOfRightInverse.apply := quotientKerAlgEquivOfRightInverse_apply -@[deprecated (since := "2024-02-27")] -alias QuotientKerAlgEquivOfRightInverseSymm.apply := quotientKerAlgEquivOfRightInverse_symm_apply - /-- The **first isomorphism theorem** for algebras. -/ @[simps!] noncomputable def quotientKerAlgEquivOfSurjective {f : A →ₐ[R₁] B} (hf : Function.Surjective f) : diff --git a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean index 0ad41dd02719a..f0bcd26091b84 100644 --- a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean @@ -183,9 +183,6 @@ lemma _root_.MvPolynomial.isHomogeneous_C_mul_X (r : R) (i : σ) : (C r * X i).IsHomogeneous 1 := (isHomogeneous_X _ _).C_mul _ -@[deprecated (since := "2024-03-21")] -alias _root_.MvPolynomial.C_mul_X := _root_.MvPolynomial.isHomogeneous_C_mul_X - lemma pow (hφ : φ.IsHomogeneous m) (n : ℕ) : (φ ^ n).IsHomogeneous (m * n) := by rw [show φ ^ n = ∏ _i ∈ Finset.range n, φ by simp] rw [show m * n = ∑ _i ∈ Finset.range n, m by simp [mul_comm]] diff --git a/Mathlib/RingTheory/PowerBasis.lean b/Mathlib/RingTheory/PowerBasis.lean index 1e67541ed7447..fc58485528703 100644 --- a/Mathlib/RingTheory/PowerBasis.lean +++ b/Mathlib/RingTheory/PowerBasis.lean @@ -75,7 +75,6 @@ theorem coe_basis (pb : PowerBasis R S) : ⇑pb.basis = fun i : Fin pb.dim => pb /-- Cannot be an instance because `PowerBasis` cannot be a class. -/ theorem finite (pb : PowerBasis R S) : Module.Finite R S := .of_basis pb.basis -@[deprecated (since := "2024-03-05")] alias finiteDimensional := PowerBasis.finite theorem finrank [StrongRankCondition R] (pb : PowerBasis R S) : Module.finrank R S = pb.dim := by diff --git a/Mathlib/RingTheory/PrincipalIdealDomain.lean b/Mathlib/RingTheory/PrincipalIdealDomain.lean index 5343a17c2253a..77a6f3b5770e5 100644 --- a/Mathlib/RingTheory/PrincipalIdealDomain.lean +++ b/Mathlib/RingTheory/PrincipalIdealDomain.lean @@ -303,12 +303,6 @@ theorem isMaximal_of_irreducible [CommRing R] [IsPrincipalIdealRing R] {p : R} refine (of_irreducible_mul hp).resolve_right (mt (fun hb => ?_) (not_le_of_lt hI)) erw [Ideal.span_singleton_le_span_singleton, IsUnit.mul_right_dvd hb]⟩⟩ -@[deprecated (since := "2024-02-12")] -protected alias irreducible_iff_prime := irreducible_iff_prime - -@[deprecated (since := "2024-02-12")] -protected alias associates_irreducible_iff_prime := associates_irreducible_iff_prime - variable [CommRing R] [IsDomain R] [IsPrincipalIdealRing R] section diff --git a/Mathlib/RingTheory/SimpleModule.lean b/Mathlib/RingTheory/SimpleModule.lean index dbae9bcb743ed..ccf7c4055b535 100644 --- a/Mathlib/RingTheory/SimpleModule.lean +++ b/Mathlib/RingTheory/SimpleModule.lean @@ -176,9 +176,6 @@ theorem IsSemisimpleModule.of_sSup_simples_eq_top (h : sSup { m : Submodule R M | IsSimpleModule R m } = ⊤) : IsSemisimpleModule R M := complementedLattice_of_sSup_atoms_eq_top (by simp_rw [← h, isSimpleModule_iff_isAtom]) -@[deprecated (since := "2024-03-05")] -alias is_semisimple_of_sSup_simples_eq_top := IsSemisimpleModule.of_sSup_simples_eq_top - namespace IsSemisimpleModule variable [IsSemisimpleModule R M] @@ -261,9 +258,6 @@ theorem sSup_simples_eq_top_iff_isSemisimpleModule : sSup { m : Submodule R M | IsSimpleModule R m } = ⊤ ↔ IsSemisimpleModule R M := ⟨.of_sSup_simples_eq_top, fun _ ↦ IsSemisimpleModule.sSup_simples_eq_top _ _⟩ -@[deprecated (since := "2024-03-05")] -alias is_semisimple_iff_top_eq_sSup_simples := sSup_simples_eq_top_iff_isSemisimpleModule - /-- A module generated by semisimple submodules is itself semisimple. -/ lemma isSemisimpleModule_of_isSemisimpleModule_submodule {s : Set ι} {p : ι → Submodule R M} (hp : ∀ i ∈ s, IsSemisimpleModule R (p i)) (hp' : ⨆ i ∈ s, p i = ⊤) : diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean index 9456594b84a19..5a7f8e147aaf6 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean @@ -135,11 +135,6 @@ instance (priority := 100) ufm_of_decomposition_of_wfDvdMonoid UniqueFactorizationMonoid α := { ‹WfDvdMonoid α› with irreducible_iff_prime := irreducible_iff_prime } -@[deprecated ufm_of_decomposition_of_wfDvdMonoid (since := "2024-02-12")] -theorem ufm_of_gcd_of_wfDvdMonoid [CancelCommMonoidWithZero α] [WfDvdMonoid α] - [DecompositionMonoid α] : UniqueFactorizationMonoid α := - ufm_of_decomposition_of_wfDvdMonoid - end Prio namespace UniqueFactorizationMonoid diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean index 2dd1d63b54515..d2a4e0e814f99 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean @@ -219,7 +219,6 @@ noncomputable def factors (a : Associates α) : FactorSet α := by theorem factors_zero : (0 : Associates α).factors = ⊤ := dif_pos rfl -@[deprecated (since := "2024-03-16")] alias factors_0 := factors_zero @[simp] theorem factors_mk (a : α) (h : a ≠ 0) : (Associates.mk a).factors = factors' a := by diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean index 9768ee91a10df..fb49eb4d8fc1f 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean @@ -128,9 +128,4 @@ theorem count_normalizedFactors_eq' [DecidableEq R] {p x : R} (hp : p = 0 ∨ Ir end multiplicity -/-- Deprecated. Use `WfDvdMonoid.max_power_factor` instead. -/ -@[deprecated WfDvdMonoid.max_power_factor (since := "2024-03-01")] -theorem max_power_factor {a₀ x : R} (h : a₀ ≠ 0) (hx : Irreducible x) : - ∃ n : ℕ, ∃ a : R, ¬x ∣ a ∧ a₀ = x ^ n * a := WfDvdMonoid.max_power_factor h hx - end UniqueFactorizationMonoid diff --git a/Mathlib/SetTheory/Cardinal/Arithmetic.lean b/Mathlib/SetTheory/Cardinal/Arithmetic.lean index 5a7416afca6f0..59acdc7b0b398 100644 --- a/Mathlib/SetTheory/Cardinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Cardinal/Arithmetic.lean @@ -463,16 +463,10 @@ theorem add_le_add_iff_of_lt_aleph0 {α β γ : Cardinal} (γ₀ : γ < ℵ₀) theorem add_nat_le_add_nat_iff {α β : Cardinal} (n : ℕ) : α + n ≤ β + n ↔ α ≤ β := add_le_add_iff_of_lt_aleph0 (nat_lt_aleph0 n) -@[deprecated (since := "2024-02-12")] -alias add_nat_le_add_nat_iff_of_lt_aleph_0 := add_nat_le_add_nat_iff - @[simp] theorem add_one_le_add_one_iff {α β : Cardinal} : α + 1 ≤ β + 1 ↔ α ≤ β := add_le_add_iff_of_lt_aleph0 one_lt_aleph0 -@[deprecated (since := "2024-02-12")] -alias add_one_le_add_one_iff_of_lt_aleph_0 := add_one_le_add_one_iff - end aleph /-! ### Properties about `power` -/ diff --git a/Mathlib/SetTheory/Cardinal/ToNat.lean b/Mathlib/SetTheory/Cardinal/ToNat.lean index 0763034be27cb..80d792d3f53b3 100644 --- a/Mathlib/SetTheory/Cardinal/ToNat.lean +++ b/Mathlib/SetTheory/Cardinal/ToNat.lean @@ -82,20 +82,9 @@ theorem toNat_lt_iff_lt_of_lt_aleph0 (hc : c < ℵ₀) (hd : d < ℵ₀) : theorem toNat_le_toNat (hcd : c ≤ d) (hd : d < ℵ₀) : toNat c ≤ toNat d := toNat_monotoneOn (hcd.trans_lt hd) hd hcd -@[deprecated toNat_le_toNat (since := "2024-02-15")] -theorem toNat_le_of_le_of_lt_aleph0 (hd : d < ℵ₀) (hcd : c ≤ d) : - toNat c ≤ toNat d := - toNat_le_toNat hcd hd - theorem toNat_lt_toNat (hcd : c < d) (hd : d < ℵ₀) : toNat c < toNat d := toNat_strictMonoOn (hcd.trans hd) hd hcd -@[deprecated toNat_lt_toNat (since := "2024-02-15")] -theorem toNat_lt_of_lt_of_lt_aleph0 (hd : d < ℵ₀) (hcd : c < d) : toNat c < toNat d := - toNat_lt_toNat hcd hd - -@[deprecated (since := "2024-02-15")] alias toNat_cast := toNat_natCast - -- See note [no_index around OfNat.ofNat] @[simp] theorem toNat_ofNat (n : ℕ) [n.AtLeastTwo] : @@ -148,11 +137,6 @@ theorem toNat_congr {β : Type v} (e : α ≃ β) : toNat #α = toNat #β := by theorem toNat_mul (x y : Cardinal) : toNat (x * y) = toNat x * toNat y := map_mul toNat x y -@[deprecated map_prod (since := "2024-02-15")] -theorem toNat_finset_prod (s : Finset α) (f : α → Cardinal) : - toNat (∏ i ∈ s, f i) = ∏ i ∈ s, toNat (f i) := - map_prod toNat _ _ - @[simp] theorem toNat_add (hc : c < ℵ₀) (hd : d < ℵ₀) : toNat (c + d) = toNat c + toNat d := by lift c to ℕ using hc @@ -164,7 +148,4 @@ theorem toNat_lift_add_lift {a : Cardinal.{u}} {b : Cardinal.{v}} (ha : a < ℵ toNat (lift.{v} a + lift.{u} b) = toNat a + toNat b := by simp [*] -@[deprecated (since := "2024-02-15")] -alias toNat_add_of_lt_aleph0 := toNat_lift_add_lift - end Cardinal diff --git a/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean b/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean index 688be889d23d8..d1448d49fe4b3 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean @@ -354,7 +354,6 @@ lemma HasProd.of_nat_of_neg_add_one {f : ℤ → M} exact (Nat.cast_injective.hasProd_range_iff.mpr hf₁).mul_isCompl this (hi₂.hasProd_range_iff.mpr hf₂) -@[deprecated (since := "2024-03-04")] alias HasSum.nonneg_add_neg := HasSum.of_nat_of_neg_add_one @[to_additive Summable.of_nat_of_neg_add_one] lemma Multipliable.of_nat_of_neg_add_one {f : ℤ → M} @@ -437,9 +436,6 @@ theorem HasProd.nat_mul_neg {f : ℤ → M} (hf : HasProd f m) : simp only [u1, u2, Nat.cast_inj, imp_self, implies_true, forall_const, prod_image, neg_inj] _ = ∏ b ∈ v', (f b * f (-b)) := prod_mul_distrib.symm⟩ -@[deprecated HasSum.nat_add_neg (since := "2024-03-04")] -alias HasSum.sum_nat_of_sum_int := HasSum.nat_add_neg - @[to_additive] theorem Multipliable.nat_mul_neg {f : ℤ → M} (hf : Multipliable f) : Multipliable fun n : ℕ ↦ f n * f (-n) := @@ -456,9 +452,6 @@ theorem HasProd.of_add_one_of_neg_add_one {f : ℤ → M} HasProd f (m * f 0 * m') := HasProd.of_nat_of_neg_add_one (mul_comm _ m ▸ HasProd.zero_mul hf₁) hf₂ -@[deprecated HasSum.of_add_one_of_neg_add_one (since := "2024-03-04")] -alias HasSum.pos_add_zero_add_neg := HasSum.of_add_one_of_neg_add_one - @[to_additive Summable.of_add_one_of_neg_add_one] lemma Multipliable.of_add_one_of_neg_add_one {f : ℤ → M} (hf₁ : Multipliable fun n : ℕ ↦ f (n + 1)) (hf₂ : Multipliable fun n : ℕ ↦ f (-(n + 1))) : @@ -490,9 +483,6 @@ lemma Multipliable.of_nat_of_neg {f : ℤ → G} (hf₁ : Multipliable fun n : (hf₂ : Multipliable fun n : ℕ ↦ f (-n)) : Multipliable f := (hf₁.hasProd.of_nat_of_neg hf₂.hasProd).multipliable -@[deprecated Summable.of_nat_of_neg (since := "2024-03-04")] -alias summable_int_of_summable_nat := Summable.of_nat_of_neg - @[to_additive] lemma tprod_of_nat_of_neg [T2Space G] {f : ℤ → G} (hf₁ : Multipliable fun n : ℕ ↦ f n) (hf₂ : Multipliable fun n : ℕ ↦ f (-n)) : diff --git a/Mathlib/Topology/Algebra/UniformRing.lean b/Mathlib/Topology/Algebra/UniformRing.lean index a52f63d8fafe8..5453277e0629f 100644 --- a/Mathlib/Topology/Algebra/UniformRing.lean +++ b/Mathlib/Topology/Algebra/UniformRing.lean @@ -233,16 +233,6 @@ theorem inseparableSetoid_ring (α) [CommRing α] [TopologicalSpace α] [Topolog Setoid.ext fun x y => addGroup_inseparable_iff.trans <| .trans (by rfl) (Submodule.quotientRel_def _).symm -@[deprecated (since := "2024-03-09")] -alias ring_sep_rel := inseparableSetoid_ring - --- Equality of types is evil -@[deprecated UniformSpace.inseparableSetoid_ring (since := "2024-02-16")] -theorem ring_sep_quot (α : Type u) [r : CommRing α] [TopologicalSpace α] [TopologicalRing α] : - SeparationQuotient α = (α ⧸ (⊥ : Ideal α).closure) := by - rw [SeparationQuotient, @inseparableSetoid_ring α r] - rfl - /-- Given a topological ring `α` equipped with a uniform structure that makes subtraction uniformly continuous, get an homeomorphism between the separated quotient of `α` and the quotient ring corresponding to the closure of zero. -/ diff --git a/Mathlib/Topology/Bases.lean b/Mathlib/Topology/Bases.lean index 691e389d9e045..db461531876af 100644 --- a/Mathlib/Topology/Bases.lean +++ b/Mathlib/Topology/Bases.lean @@ -526,15 +526,9 @@ theorem isSeparable_range [TopologicalSpace β] [SeparableSpace α] {f : α → theorem IsSeparable.of_subtype (s : Set α) [SeparableSpace s] : IsSeparable s := by simpa using isSeparable_range (continuous_subtype_val (p := (· ∈ s))) -@[deprecated (since := "2024-02-05")] -alias isSeparable_of_separableSpace_subtype := IsSeparable.of_subtype - theorem IsSeparable.of_separableSpace [h : SeparableSpace α] (s : Set α) : IsSeparable s := IsSeparable.mono (isSeparable_univ_iff.2 h) (subset_univ _) -@[deprecated (since := "2024-02-05")] -alias isSeparable_of_separableSpace := IsSeparable.of_separableSpace - end TopologicalSpace open TopologicalSpace diff --git a/Mathlib/Topology/CompactOpen.lean b/Mathlib/Topology/CompactOpen.lean index 49f2b425934f9..3f9c290325a74 100644 --- a/Mathlib/Topology/CompactOpen.lean +++ b/Mathlib/Topology/CompactOpen.lean @@ -279,17 +279,11 @@ theorem compactOpen_eq_iInf_induced : convert isOpen_induced (isOpen_setOf_mapsTo (isCompact_iff_isCompact_univ.1 hK) hU) simp [mapsTo_univ_iff, Subtype.forall, MapsTo] -@[deprecated (since := "2024-03-05")] -alias compactOpen_eq_sInf_induced := compactOpen_eq_iInf_induced - theorem nhds_compactOpen_eq_iInf_nhds_induced (f : C(X, Y)) : 𝓝 f = ⨅ (s) (_ : IsCompact s), (𝓝 (f.restrict s)).comap (ContinuousMap.restrict s) := by rw [compactOpen_eq_iInf_induced] simp only [nhds_iInf, nhds_induced] -@[deprecated (since := "2024-03-05")] -alias nhds_compactOpen_eq_sInf_nhds_induced := nhds_compactOpen_eq_iInf_nhds_induced - theorem tendsto_compactOpen_restrict {ι : Type*} {l : Filter ι} {F : ι → C(X, Y)} {f : C(X, Y)} (hFf : Filter.Tendsto F l (𝓝 f)) (s : Set X) : Tendsto (fun i => (F i).restrict s) l (𝓝 (f.restrict s)) := @@ -371,15 +365,6 @@ def curry (f : C(X × Y, Z)) : C(X, C(Y, Z)) where theorem curry_apply (f : C(X × Y, Z)) (a : X) (b : Y) : f.curry a b = f (a, b) := rfl -/-- Auxiliary definition, see `ContinuousMap.curry` and `Homeomorph.curry`. -/ -@[deprecated ContinuousMap.curry (since := "2024-03-05")] -def curry' (f : C(X × Y, Z)) (a : X) : C(Y, Z) := curry f a - -set_option linter.deprecated false in -/-- If a map `α × β → γ` is continuous, then its curried form `α → C(β, γ)` is continuous. -/ -@[deprecated ContinuousMap.curry (since := "2024-03-05")] -theorem continuous_curry' (f : C(X × Y, Z)) : Continuous (curry' f) := (curry f).continuous - /-- To show continuity of a map `α → C(β, γ)`, it suffices to show that its uncurried form `α × β → γ` is continuous. -/ theorem continuous_of_continuous_uncurry (f : X → C(Y, Z)) diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index 15e1b303bf533..3482d2379b813 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -283,10 +283,6 @@ theorem IsCompact.nonempty_iInter_of_directed_nonempty_isCompact_isClosed rcases htd i₀ i with ⟨j, hji₀, hji⟩ exact (htn j).mono (subset_inter hji₀ hji) -@[deprecated (since := "2024-02-28")] -alias IsCompact.nonempty_iInter_of_directed_nonempty_compact_closed := - IsCompact.nonempty_iInter_of_directed_nonempty_isCompact_isClosed - /-- Cantor's intersection theorem for `sInter`: the intersection of a directed family of nonempty compact closed sets is nonempty. -/ theorem IsCompact.nonempty_sInter_of_directed_nonempty_isCompact_isClosed @@ -307,10 +303,6 @@ theorem IsCompact.nonempty_iInter_of_sequence_nonempty_isCompact_isClosed (t : have htc : ∀ i, IsCompact (t i) := fun i => ht0.of_isClosed_subset (htcl i) (this i) IsCompact.nonempty_iInter_of_directed_nonempty_isCompact_isClosed t htd htn htc htcl -@[deprecated (since := "2024-02-28")] -alias IsCompact.nonempty_iInter_of_sequence_nonempty_compact_closed := - IsCompact.nonempty_iInter_of_sequence_nonempty_isCompact_isClosed - /-- For every open cover of a compact set, there exists a finite subcover. -/ theorem IsCompact.elim_finite_subcover_image {b : Set ι} {c : ι → Set X} (hs : IsCompact s) (hc₁ : ∀ i ∈ b, IsOpen (c i)) (hc₂ : s ⊆ ⋃ i ∈ b, c i) : @@ -640,14 +632,6 @@ theorem mem_coclosedCompact_iff : closure_minimal (compl_subset_comm.2 hst) htcl · exact ⟨closure sᶜ, ⟨isClosed_closure, h⟩, compl_subset_comm.2 subset_closure⟩ -@[deprecated mem_coclosedCompact_iff (since := "2024-02-16")] -theorem mem_coclosedCompact : s ∈ coclosedCompact X ↔ ∃ t, IsClosed t ∧ IsCompact t ∧ tᶜ ⊆ s := by - simp only [hasBasis_coclosedCompact.mem_iff, and_assoc] - -@[deprecated mem_coclosedCompact_iff (since := "2024-02-16")] -theorem mem_coclosed_compact' : s ∈ coclosedCompact X ↔ ∃ t, IsClosed t ∧ IsCompact t ∧ sᶜ ⊆ t := by - simp only [hasBasis_coclosedCompact.mem_iff, compl_subset_comm, and_assoc] - /-- Complement of a set belongs to `coclosedCompact` if and only if its closure is compact. -/ theorem compl_mem_coclosedCompact : sᶜ ∈ coclosedCompact X ↔ IsCompact (closure s) := by rw [mem_coclosedCompact_iff, compl_compl] diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index ee4d5a39eb471..5e5987a52ca6d 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -467,8 +467,6 @@ theorem Continuous.uncurry_right {f : X → Y → Z} (y : Y) (h : Continuous (un Continuous fun a => f a y := h.comp (Continuous.Prod.mk_left _) -@[deprecated (since := "2024-03-09")] alias continuous_uncurry_left := Continuous.uncurry_left -@[deprecated (since := "2024-03-09")] alias continuous_uncurry_right := Continuous.uncurry_right theorem continuous_curry {g : X × Y → Z} (x : X) (h : Continuous g) : Continuous (curry g x) := Continuous.uncurry_left x h diff --git a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean index c634039f9359d..0d56fcaf0d007 100644 --- a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean @@ -1414,8 +1414,6 @@ instance instLattice : Lattice (α →ᵇ β) := DFunLike.coe_injective.lattice @[simp, norm_cast] lemma coe_posPart (f : α →ᵇ β) : ⇑f⁺ = (⇑f)⁺ := rfl @[simp, norm_cast] lemma coe_negPart (f : α →ᵇ β) : ⇑f⁻ = (⇑f)⁻ := rfl -@[deprecated (since := "2024-02-21")] alias coeFn_sup := coe_sup -@[deprecated (since := "2024-02-21")] alias coeFn_abs := coe_abs instance instNormedLatticeAddCommGroup : NormedLatticeAddCommGroup (α →ᵇ β) := { instSeminormedAddCommGroup with diff --git a/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean b/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean index eecb451c02ae7..6726cd76211d7 100644 --- a/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean +++ b/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean @@ -430,7 +430,6 @@ theorem isClosed_range_toBCF : IsClosed (range (toBCF : C₀(α, β) → α → _ < ε := by simpa [add_halves ε] using add_lt_add_right (mem_ball.1 hg) (ε / 2) exact ⟨⟨f.toContinuousMap, this⟩, rfl⟩ -@[deprecated (since := "2024-03-17")] alias closed_range_toBCF := isClosed_range_toBCF /-- Continuous functions vanishing at infinity taking values in a complete space form a complete space. -/ diff --git a/Mathlib/Topology/GDelta/Basic.lean b/Mathlib/Topology/GDelta/Basic.lean index 60e2975ce2246..250e1184d4569 100644 --- a/Mathlib/Topology/GDelta/Basic.lean +++ b/Mathlib/Topology/GDelta/Basic.lean @@ -66,25 +66,21 @@ theorem IsOpen.isGδ {s : Set X} (h : IsOpen s) : IsGδ s := protected theorem IsGδ.empty : IsGδ (∅ : Set X) := isOpen_empty.isGδ -@[deprecated (since := "2024-02-15")] alias isGδ_empty := IsGδ.empty @[simp] protected theorem IsGδ.univ : IsGδ (univ : Set X) := isOpen_univ.isGδ -@[deprecated (since := "2024-02-15")] alias isGδ_univ := IsGδ.univ theorem IsGδ.biInter_of_isOpen {I : Set ι} (hI : I.Countable) {f : ι → Set X} (hf : ∀ i ∈ I, IsOpen (f i)) : IsGδ (⋂ i ∈ I, f i) := ⟨f '' I, by rwa [forall_mem_image], hI.image _, by rw [sInter_image]⟩ -@[deprecated (since := "2024-02-15")] alias isGδ_biInter_of_isOpen := IsGδ.biInter_of_isOpen theorem IsGδ.iInter_of_isOpen [Countable ι'] {f : ι' → Set X} (hf : ∀ i, IsOpen (f i)) : IsGδ (⋂ i, f i) := ⟨range f, by rwa [forall_mem_range], countable_range _, by rw [sInter_range]⟩ -@[deprecated (since := "2024-02-15")] alias isGδ_iInter_of_isOpen := IsGδ.iInter_of_isOpen lemma isGδ_iff_eq_iInter_nat {s : Set X} : IsGδ s ↔ ∃ (f : ℕ → Set X), (∀ n, IsOpen (f n)) ∧ s = ⋂ n, f n := by @@ -115,13 +111,11 @@ theorem IsGδ.biInter {s : Set ι} (hs : s.Countable) {t : ∀ i ∈ s, Set X} haveI := hs.to_subtype exact .iInter fun x => ht x x.2 -@[deprecated (since := "2024-02-15")] alias isGδ_biInter := IsGδ.biInter /-- A countable intersection of Gδ sets is a Gδ set. -/ theorem IsGδ.sInter {S : Set (Set X)} (h : ∀ s ∈ S, IsGδ s) (hS : S.Countable) : IsGδ (⋂₀ S) := by simpa only [sInter_eq_biInter] using IsGδ.biInter hS h -@[deprecated (since := "2024-02-15")] alias isGδ_sInter := IsGδ.sInter theorem IsGδ.inter {s t : Set X} (hs : IsGδ s) (ht : IsGδ t) : IsGδ (s ∩ t) := by rw [inter_eq_iInter] @@ -150,9 +144,6 @@ theorem IsGδ.biUnion {s : Set ι} (hs : s.Finite) {f : ι → Set X} (h : ∀ i rw [← sUnion_image] exact .sUnion (hs.image _) (forall_mem_image.2 h) -@[deprecated (since := "2024-02-15")] -alias isGδ_biUnion := IsGδ.biUnion - /-- The union of finitely many Gδ sets is a Gδ set, bounded indexed union version. -/ theorem IsGδ.iUnion [Finite ι'] {f : ι' → Set X} (h : ∀ i, IsGδ (f i)) : IsGδ (⋃ i, f i) := .sUnion (finite_range _) <| forall_mem_range.2 h diff --git a/Mathlib/Topology/GDelta/UniformSpace.lean b/Mathlib/Topology/GDelta/UniformSpace.lean index 3aab2f9dd00c0..735d3d2bd0269 100644 --- a/Mathlib/Topology/GDelta/UniformSpace.lean +++ b/Mathlib/Topology/GDelta/UniformSpace.lean @@ -50,6 +50,5 @@ theorem IsGδ.setOf_continuousAt [UniformSpace Y] [IsCountablyGenerated (𝓤 Y) rintro ⟨s, ⟨hsx, hso⟩, hsU⟩ filter_upwards [IsOpen.mem_nhds hso hsx] with _ hy using ⟨s, ⟨hy, hso⟩, hsU⟩ -@[deprecated (since := "2024-02-15")] alias isGδ_setOf_continuousAt := IsGδ.setOf_continuousAt end ContinuousAt diff --git a/Mathlib/Topology/Instances/Discrete.lean b/Mathlib/Topology/Instances/Discrete.lean index b9af3e7060b0e..d798a8a6591b8 100644 --- a/Mathlib/Topology/Instances/Discrete.lean +++ b/Mathlib/Topology/Instances/Discrete.lean @@ -36,11 +36,6 @@ instance (priority := 100) DiscreteTopology.secondCountableTopology_of_countable secondCountableTopology_of_countable_cover (singletons_open_iff_discrete.mpr hd) (iUnion_of_singleton α) -@[deprecated DiscreteTopology.secondCountableTopology_of_countable (since := "2024-03-11")] -theorem DiscreteTopology.secondCountableTopology_of_encodable {α : Type*} - [TopologicalSpace α] [DiscreteTopology α] [Countable α] : SecondCountableTopology α := - DiscreteTopology.secondCountableTopology_of_countable - theorem LinearOrder.bot_topologicalSpace_eq_generateFrom {α} [LinearOrder α] [PredOrder α] [SuccOrder α] : (⊥ : TopologicalSpace α) = generateFrom { s | ∃ a, s = Ioi a ∨ s = Iio a } := by let _ := Preorder.topology α diff --git a/Mathlib/Topology/Instances/Int.lean b/Mathlib/Topology/Instances/Int.lean index 009674177e1f0..ba41c15cda783 100644 --- a/Mathlib/Topology/Instances/Int.lean +++ b/Mathlib/Topology/Instances/Int.lean @@ -73,7 +73,6 @@ theorem cobounded_eq : Bornology.cobounded ℤ = atBot ⊔ atTop := by simp_rw [← comap_dist_right_atTop (0 : ℤ), dist_eq', sub_zero, ← comap_abs_atTop, ← @Int.comap_cast_atTop ℝ, comap_comap]; rfl -@[deprecated (since := "2024-02-07")] alias cocompact_eq := cocompact_eq_atBot_atTop @[simp] theorem cofinite_eq : (cofinite : Filter ℤ) = atBot ⊔ atTop := by diff --git a/Mathlib/Topology/Instances/Irrational.lean b/Mathlib/Topology/Instances/Irrational.lean index f395be578fdff..8b82916e9b93b 100644 --- a/Mathlib/Topology/Instances/Irrational.lean +++ b/Mathlib/Topology/Instances/Irrational.lean @@ -36,7 +36,6 @@ open Filter Topology protected theorem IsGδ.setOf_irrational : IsGδ { x | Irrational x } := (countable_range _).isGδ_compl -@[deprecated (since := "2024-02-15")] alias isGδ_irrational := IsGδ.setOf_irrational theorem dense_irrational : Dense { x : ℝ | Irrational x } := by refine Real.isTopologicalBasis_Ioo_rat.dense_iff.2 ?_ diff --git a/Mathlib/Topology/Instances/Real.lean b/Mathlib/Topology/Instances/Real.lean index 522b40e629e4b..41f5e2e54a206 100644 --- a/Mathlib/Topology/Instances/Real.lean +++ b/Mathlib/Topology/Instances/Real.lean @@ -74,10 +74,7 @@ theorem Real.isTopologicalBasis_Ioo_rat : theorem Real.cobounded_eq : cobounded ℝ = atBot ⊔ atTop := by simp only [← comap_dist_right_atTop (0 : ℝ), Real.dist_eq, sub_zero, comap_abs_atTop] -@[deprecated (since := "2024-02-07")] alias Real.cocompact_eq := cocompact_eq_atBot_atTop -@[deprecated (since := "2024-02-07")] alias Real.atBot_le_cocompact := atBot_le_cocompact -@[deprecated (since := "2024-02-07")] alias Real.atTop_le_cocompact := atTop_le_cocompact /- TODO(Mario): Prove that these are uniform isomorphisms instead of uniform embeddings lemma uniform_embedding_add_rat {r : ℚ} : uniform_embedding (fun p : ℚ => p + r) := diff --git a/Mathlib/Topology/Maps/Basic.lean b/Mathlib/Topology/Maps/Basic.lean index 077466f86f877..ecd68d48bac01 100644 --- a/Mathlib/Topology/Maps/Basic.lean +++ b/Mathlib/Topology/Maps/Basic.lean @@ -487,7 +487,6 @@ theorem of_nonempty (h : ∀ s, IsClosed s → s.Nonempty → IsClosed (f '' s)) theorem isClosed_range (hf : IsClosedMap f) : IsClosed (range f) := @image_univ _ _ f ▸ hf _ isClosed_univ -@[deprecated (since := "2024-03-17")] alias closed_range := isClosed_range theorem isQuotientMap (hcl : IsClosedMap f) (hcont : Continuous f) (hsurj : Surjective f) : IsQuotientMap f := diff --git a/Mathlib/Topology/MetricSpace/Completion.lean b/Mathlib/Topology/MetricSpace/Completion.lean index cbca7d20888b3..69037d369df59 100644 --- a/Mathlib/Topology/MetricSpace/Completion.lean +++ b/Mathlib/Topology/MetricSpace/Completion.lean @@ -158,10 +158,6 @@ instance instMetricSpace : MetricSpace (Completion α) := toUniformSpace := inferInstance uniformity_dist := Completion.uniformity_dist } _ -@[deprecated eq_of_dist_eq_zero (since := "2024-03-10")] -protected theorem eq_of_dist_eq_zero (x y : Completion α) (h : dist x y = 0) : x = y := - eq_of_dist_eq_zero h - /-- The embedding of a metric space in its completion is an isometry. -/ theorem coe_isometry : Isometry ((↑) : α → Completion α) := Isometry.of_dist_eq Completion.dist_eq diff --git a/Mathlib/Topology/MetricSpace/Polish.lean b/Mathlib/Topology/MetricSpace/Polish.lean index 70ae87a47d620..543d4da8c0d76 100644 --- a/Mathlib/Topology/MetricSpace/Polish.lean +++ b/Mathlib/Topology/MetricSpace/Polish.lean @@ -102,9 +102,6 @@ instance (priority := 100) instMetrizableSpace (α : Type*) [TopologicalSpace α letI := upgradePolishSpace α infer_instance -@[deprecated "No deprecation message was provided." (since := "2024-02-23")] -theorem t2Space (α : Type*) [TopologicalSpace α] [PolishSpace α] : T2Space α := inferInstance - /-- A countable product of Polish spaces is Polish. -/ instance pi_countable {ι : Type*} [Countable ι] {E : ι → Type*} [∀ i, TopologicalSpace (E i)] [∀ i, PolishSpace (E i)] : PolishSpace (∀ i, E i) := by diff --git a/Mathlib/Topology/MetricSpace/Sequences.lean b/Mathlib/Topology/MetricSpace/Sequences.lean index ed02f21cdebd7..5548bfac6db34 100644 --- a/Mathlib/Topology/MetricSpace/Sequences.lean +++ b/Mathlib/Topology/MetricSpace/Sequences.lean @@ -17,12 +17,6 @@ open scoped Topology variable {X : Type*} [PseudoMetricSpace X] -@[deprecated lebesgue_number_lemma_of_metric (since := "2024-02-24")] -nonrec theorem SeqCompact.lebesgue_number_lemma_of_metric {ι : Sort*} {c : ι → Set X} {s : Set X} - (hs : IsSeqCompact s) (hc₁ : ∀ i, IsOpen (c i)) (hc₂ : s ⊆ ⋃ i, c i) : - ∃ δ > 0, ∀ a ∈ s, ∃ i, ball a δ ⊆ c i := - lebesgue_number_lemma_of_metric hs.isCompact hc₁ hc₂ - variable [ProperSpace X] {s : Set X} /-- A version of **Bolzano-Weierstrass**: in a proper metric space (eg. $ℝ^n$), diff --git a/Mathlib/Topology/Order.lean b/Mathlib/Topology/Order.lean index 0be6eb00b68ce..2bd638f106979 100644 --- a/Mathlib/Topology/Order.lean +++ b/Mathlib/Topology/Order.lean @@ -286,9 +286,6 @@ theorem le_of_nhds_le_nhds (h : ∀ x, @nhds α t₁ x ≤ @nhds α t₂ x) : t rw [@isOpen_iff_mem_nhds _ _ t₁, @isOpen_iff_mem_nhds α _ t₂] exact fun hs a ha => h _ (hs _ ha) -@[deprecated (since := "2024-03-01")] -alias eq_of_nhds_eq_nhds := TopologicalSpace.ext_nhds - theorem eq_bot_of_singletons_open {t : TopologicalSpace α} (h : ∀ x, IsOpen[t] {x}) : t = ⊥ := bot_unique fun s _ => biUnion_of_singleton s ▸ isOpen_biUnion fun x _ => h x @@ -572,19 +569,11 @@ theorem nhds_nhdsAdjoint_same (a : α) (f : Filter α) : exact IsOpen.mem_nhds (fun _ ↦ htf) hat · exact sup_le (pure_le_nhds _) ((gc_nhds a).le_u_l f) -@[deprecated (since := "2024-02-10")] -alias nhdsAdjoint_nhds := nhds_nhdsAdjoint_same - theorem nhds_nhdsAdjoint_of_ne {a b : α} (f : Filter α) (h : b ≠ a) : @nhds α (nhdsAdjoint a f) b = pure b := let _ := nhdsAdjoint a f (isOpen_singleton_iff_nhds_eq_pure _).1 <| isOpen_singleton_nhdsAdjoint f h -@[deprecated nhds_nhdsAdjoint_of_ne (since := "2024-02-10")] -theorem nhdsAdjoint_nhds_of_ne (a : α) (f : Filter α) {b : α} (h : b ≠ a) : - @nhds α (nhdsAdjoint a f) b = pure b := - nhds_nhdsAdjoint_of_ne f h - theorem nhds_nhdsAdjoint [DecidableEq α] (a : α) (f : Filter α) : @nhds α (nhdsAdjoint a f) = update pure a (pure a ⊔ f) := eq_update_iff.2 ⟨nhds_nhdsAdjoint_same .., fun _ ↦ nhds_nhdsAdjoint_of_ne _⟩ diff --git a/Mathlib/Topology/Order/Basic.lean b/Mathlib/Topology/Order/Basic.lean index a0884956957dd..80dc91a3e3e3e 100644 --- a/Mathlib/Topology/Order/Basic.lean +++ b/Mathlib/Topology/Order/Basic.lean @@ -489,16 +489,6 @@ theorem Dense.topology_eq_generateFrom [OrderTopology α] [DenselyOrdered α] {s let _ := generateFrom (Ioi '' s ∪ Iio '' s) exact isOpen_iUnion fun x ↦ isOpen_iUnion fun h ↦ .basic _ <| .inr <| mem_image_of_mem _ h.1 -@[deprecated OrderBot.atBot_eq (since := "2024-02-14")] -theorem atBot_le_nhds_bot [OrderBot α] : (atBot : Filter α) ≤ 𝓝 ⊥ := by - rw [OrderBot.atBot_eq] - apply pure_le_nhds - -set_option linter.deprecated false in -@[deprecated OrderTop.atTop_eq (since := "2024-02-14")] -theorem atTop_le_nhds_top [OrderTop α] : (atTop : Filter α) ≤ 𝓝 ⊤ := - @atBot_le_nhds_bot αᵒᵈ _ _ _ - variable (α) /-- Let `α` be a densely ordered linear order with order topology. If `α` is a separable space, then diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index 94bdce684cd4a..846125f18cf39 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -108,10 +108,6 @@ variable [TopologicalSpace α] [Preorder α] [ClosedIicTopology α] {f : β → theorem isClosed_Iic : IsClosed (Iic a) := ClosedIicTopology.isClosed_Iic a -@[deprecated isClosed_Iic (since := "2024-02-15")] -lemma ClosedIicTopology.isClosed_le' (a : α) : IsClosed {x | x ≤ a} := isClosed_Iic a -export ClosedIicTopology (isClosed_le') - instance : ClosedIciTopology αᵒᵈ where isClosed_Ici _ := isClosed_Iic (α := α) @@ -415,10 +411,6 @@ variable [TopologicalSpace α] [Preorder α] [ClosedIciTopology α] {f : β → theorem isClosed_Ici {a : α} : IsClosed (Ici a) := ClosedIciTopology.isClosed_Ici a -@[deprecated isClosed_Ici (since := "2024-02-15")] -lemma ClosedIciTopology.isClosed_ge' (a : α) : IsClosed {x | a ≤ x} := isClosed_Ici a -export ClosedIciTopology (isClosed_ge') - instance : ClosedIicTopology αᵒᵈ where isClosed_Iic _ := isClosed_Ici (α := α) diff --git a/Mathlib/Topology/Semicontinuous.lean b/Mathlib/Topology/Semicontinuous.lean index ae630cd687ba7..b569bcd00df4f 100644 --- a/Mathlib/Topology/Semicontinuous.lean +++ b/Mathlib/Topology/Semicontinuous.lean @@ -326,14 +326,8 @@ theorem lowerSemicontinuous_iff_isClosed_epigraph {f : α → γ} : · rw [lowerSemicontinuous_iff_isClosed_preimage] exact fun hf y ↦ hf.preimage (Continuous.Prod.mk_left y) -@[deprecated (since := "2024-03-02")] -alias lowerSemicontinuous_iff_IsClosed_epigraph := lowerSemicontinuous_iff_isClosed_epigraph - alias ⟨LowerSemicontinuous.isClosed_epigraph, _⟩ := lowerSemicontinuous_iff_isClosed_epigraph -@[deprecated (since := "2024-03-02")] -alias LowerSemicontinuous.IsClosed_epigraph := LowerSemicontinuous.isClosed_epigraph - end /-! ### Composition -/ diff --git a/Mathlib/Topology/Separation/GDelta.lean b/Mathlib/Topology/Separation/GDelta.lean index abc821f17ade6..57bc39f4a2d32 100644 --- a/Mathlib/Topology/Separation/GDelta.lean +++ b/Mathlib/Topology/Separation/GDelta.lean @@ -35,7 +35,6 @@ section Separation theorem IsGδ.compl_singleton (x : X) [T1Space X] : IsGδ ({x}ᶜ : Set X) := isOpen_compl_singleton.isGδ -@[deprecated (since := "2024-02-15")] alias isGδ_compl_singleton := IsGδ.compl_singleton theorem Set.Countable.isGδ_compl {s : Set X} [T1Space X] (hs : s.Countable) : IsGδ sᶜ := by rw [← biUnion_of_singleton s, compl_iUnion₂] @@ -56,7 +55,6 @@ protected theorem IsGδ.singleton [FirstCountableTopology X] [T1Space X] (x : X) rw [← biInter_basis_nhds h_basis.toHasBasis] exact .biInter (to_countable _) fun n _ => (hU n).2.isGδ -@[deprecated (since := "2024-02-15")] alias isGδ_singleton := IsGδ.singleton theorem Set.Finite.isGδ [FirstCountableTopology X] {s : Set X} [T1Space X] (hs : s.Finite) : IsGδ s := diff --git a/Mathlib/Topology/Separation/Hausdorff.lean b/Mathlib/Topology/Separation/Hausdorff.lean index 17cc8d6659b95..957b2b4b1cfd1 100644 --- a/Mathlib/Topology/Separation/Hausdorff.lean +++ b/Mathlib/Topology/Separation/Hausdorff.lean @@ -528,9 +528,6 @@ theorem Function.LeftInverse.isClosed_range [T2Space X] {f : X → Y} {g : Y → h.rightInvOn_range.eqOn.closure (hg.comp hf) continuous_id isClosed_of_closure_subset fun x hx => ⟨f x, this hx⟩ -@[deprecated (since := "2024-03-17")] -alias Function.LeftInverse.closed_range := Function.LeftInverse.isClosed_range - theorem Function.LeftInverse.isClosedEmbedding [T2Space X] {f : X → Y} {g : Y → X} (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) : IsClosedEmbedding g := ⟨.of_leftInverse h hf hg, h.isClosed_range hf hg⟩ diff --git a/Mathlib/Topology/Separation/Regular.lean b/Mathlib/Topology/Separation/Regular.lean index f886e9d275213..3be8602edb54b 100644 --- a/Mathlib/Topology/Separation/Regular.lean +++ b/Mathlib/Topology/Separation/Regular.lean @@ -113,24 +113,15 @@ theorem RegularSpace.of_lift'_closure_le (h : ∀ x : X, (𝓝 x).lift' closure theorem RegularSpace.of_lift'_closure (h : ∀ x : X, (𝓝 x).lift' closure = 𝓝 x) : RegularSpace X := Iff.mpr ((regularSpace_TFAE X).out 0 5) h -@[deprecated (since := "2024-02-28")] -alias RegularSpace.ofLift'_closure := RegularSpace.of_lift'_closure - theorem RegularSpace.of_hasBasis {ι : X → Sort*} {p : ∀ a, ι a → Prop} {s : ∀ a, ι a → Set X} (h₁ : ∀ a, (𝓝 a).HasBasis (p a) (s a)) (h₂ : ∀ a i, p a i → IsClosed (s a i)) : RegularSpace X := .of_lift'_closure fun a => (h₁ a).lift'_closure_eq_self (h₂ a) -@[deprecated (since := "2024-02-28")] -alias RegularSpace.ofBasis := RegularSpace.of_hasBasis - theorem RegularSpace.of_exists_mem_nhds_isClosed_subset (h : ∀ (x : X), ∀ s ∈ 𝓝 x, ∃ t ∈ 𝓝 x, IsClosed t ∧ t ⊆ s) : RegularSpace X := Iff.mpr ((regularSpace_TFAE X).out 0 3) h -@[deprecated (since := "2024-02-28")] -alias RegularSpace.ofExistsMemNhdsIsClosedSubset := RegularSpace.of_exists_mem_nhds_isClosed_subset - /-- A weakly locally compact R₁ space is regular. -/ instance (priority := 100) [WeaklyLocallyCompactSpace X] [R1Space X] : RegularSpace X := .of_hasBasis isCompact_isClosed_basis_nhds fun _ _ ⟨_, _, h⟩ ↦ h diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index 0057d6c5ff78f..f5abe4496c48a 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -352,20 +352,6 @@ theorem UniformSpace.toCore_toTopologicalSpace (u : UniformSpace α) : TopologicalSpace.ext_nhds fun a ↦ by rw [u.nhds_eq_comap_uniformity, u.toCore.nhds_toTopologicalSpace] -/-- Build a `UniformSpace` from a `UniformSpace.Core` and a compatible topology. -Use `UniformSpace.mk` instead to avoid proving -the unnecessary assumption `UniformSpace.Core.refl`. - -The main constructor used to use a different compatibility assumption. -This definition was created as a step towards porting to a new definition. -Now the main definition is ported, -so this constructor will be removed in a few months. -/ -@[deprecated UniformSpace.mk (since := "2024-03-20")] -def UniformSpace.ofNhdsEqComap (u : UniformSpace.Core α) (_t : TopologicalSpace α) - (h : ∀ x, 𝓝 x = u.uniformity.comap (Prod.mk x)) : UniformSpace α where - __ := u - nhds_eq_comap_uniformity := h - @[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 diff --git a/Mathlib/Topology/UnitInterval.lean b/Mathlib/Topology/UnitInterval.lean index cec95ceaeefe6..5cecb700fae18 100644 --- a/Mathlib/Topology/UnitInterval.lean +++ b/Mathlib/Topology/UnitInterval.lean @@ -122,8 +122,6 @@ def symmHomeomorph : I ≃ₜ I where theorem strictAnti_symm : StrictAnti σ := fun _ _ h ↦ sub_lt_sub_left (α := ℝ) h _ -@[deprecated (since := "2024-02-27")] alias involutive_symm := symm_involutive -@[deprecated (since := "2024-02-27")] alias bijective_symm := symm_bijective @[simp] theorem symm_inj {i j : I} : σ i = σ j ↔ i = j := symm_bijective.injective.eq_iff From 3279a44d606aac10f2909dd10d0ece578c78bf2f Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Mon, 6 Jan 2025 11:06:13 +0000 Subject: [PATCH 173/189] feat(CategoryTheory): products in the category of Ind-objects (#20507) Follow-up to #18988. --- .../Limits/Indization/Category.lean | 23 ++++++++++++++++++- 1 file changed, 22 insertions(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Limits/Indization/Category.lean b/Mathlib/CategoryTheory/Limits/Indization/Category.lean index c927eac12564d..4de4c45fb1b0f 100644 --- a/Mathlib/CategoryTheory/Limits/Indization/Category.lean +++ b/Mathlib/CategoryTheory/Limits/Indization/Category.lean @@ -8,6 +8,7 @@ import Mathlib.CategoryTheory.Limits.Constructions.Filtered import Mathlib.CategoryTheory.Limits.FullSubcategory import Mathlib.CategoryTheory.Limits.Indization.LocallySmall import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits +import Mathlib.CategoryTheory.Limits.Indization.Products /-! # The category of Ind-objects @@ -24,7 +25,10 @@ Adopting the theorem numbering of [Kashiwara2006], we show that * `C ⥤ Ind C` preserves finite colimits (6.1.6), * if `C` has coproducts indexed by a finite type `α`, then `Ind C` has coproducts indexed by `α` (6.1.18(ii)), -* if `C` has finite coproducts, then `Ind C` has small coproducts (6.1.18(ii)). +* if `C` has finite coproducts, then `Ind C` has small coproducts (6.1.18(ii)), +* if `C` has products indexed by `α`, then `Ind C` has products indexed by `α`, and the functor + `Ind C ⥤ Cᵒᵖ ⥤ Type v` creates such products (6.1.17) +* the functor `C ⥤ Ind C` preserves small limits. More limit-colimit properties will follow. @@ -107,6 +111,23 @@ instance : HasFilteredColimits (Ind C) where HasColimitsOfShape _ _ _ := hasColimitsOfShape_of_hasColimitsOfShape_createsColimitsOfShape (Ind.inclusion C) +noncomputable instance {J : Type v} [HasLimitsOfShape (Discrete J) C] : + CreatesLimitsOfShape (Discrete J) (Ind.inclusion C) := + letI _ : CreatesLimitsOfShape (Discrete J) (fullSubcategoryInclusion (IsIndObject (C := C))) := + createsLimitsOfShapeFullSubcategoryInclusion (closedUnderLimitsOfShape_of_limit + (isIndObject_limit_of_discrete_of_hasLimitsOfShape _)) + inferInstanceAs <| + CreatesLimitsOfShape (Discrete J) ((Ind.equivalence C).functor ⋙ fullSubcategoryInclusion _) + +instance {J : Type v} [HasLimitsOfShape (Discrete J) C] : + HasLimitsOfShape (Discrete J) (Ind C) := + hasLimitsOfShape_of_hasLimitsOfShape_createsLimitsOfShape (Ind.inclusion C) + +instance : PreservesLimits (Ind.yoneda (C := C)) := + letI _ : PreservesLimitsOfSize.{v, v} (Ind.yoneda ⋙ Ind.inclusion C) := + preservesLimits_of_natIso Ind.yonedaCompInclusion.symm + preservesLimits_of_reflects_of_preserves Ind.yoneda (Ind.inclusion C) + theorem Ind.isIndObject_inclusion_obj (X : Ind C) : IsIndObject ((Ind.inclusion C).obj X) := X.2 From 14950c530f22c7bac6bee3168005c16dce3485f2 Mon Sep 17 00:00:00 2001 From: smorel394 Date: Mon, 6 Jan 2025 11:06:14 +0000 Subject: [PATCH 174/189] refactor(CategoryTheory/Limits/Preserves/Ulift): simplify the proof that the universe lifting functor for types preserves all colimits (#20508) The current proof that the universe lifting functor for types preserves all colimits is redundant with the work done in `CategoryTheory.Limits.Types`, in particular with `isColimit_iff_bijective_desc`, as the result actually follows almost immediately from this last statement. This PR simplifies the proof using `isColimit_iff_bijective_desc`. Co-authored-by: smorel394 <67864981+smorel394@users.noreply.github.com> --- .../Limits/Preserves/Ulift.lean | 99 +++---------------- Mathlib/CategoryTheory/Limits/Types.lean | 52 +++++++++- 2 files changed, 62 insertions(+), 89 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean b/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean index d2998c6990db4..2b845755423e4 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Dagur Asgeirsson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Dagur Asgeirsson, Junyan Xu +Authors: Dagur Asgeirsson, Junyan Xu, Sophie Morel -/ import Mathlib.CategoryTheory.Limits.Creates import Mathlib.CategoryTheory.Limits.Types @@ -10,11 +10,12 @@ import Mathlib.Data.Set.Subsingleton /-! # `ULift` creates small (co)limits -This file shows that `uliftFunctor.{v, u}` creates (and hence preserves) limits and colimits indexed -by `J` with `[Category.{w, u} J]`. -It also shows that `uliftFunctor.{v, u}` *preserves* "all" limits, -potentially too big to exist in `Type u`). +This file shows that `uliftFunctor.{v, u}` preserves all limits and colimits, including those +potentially too big to exist in `Type u`. + +As this functor is fully faithful, we also deduce that it creates `u`-small limits and +colimits. -/ @@ -37,8 +38,7 @@ def sectionsEquiv {J : Type*} [Category J] (K : J ⥤ Type u) : /-- The functor `uliftFunctor : Type u ⥤ Type (max u v)` preserves limits of arbitrary size. -/ -noncomputable -instance : PreservesLimitsOfSize.{w', w} uliftFunctor.{v, u} where +noncomputable instance : PreservesLimitsOfSize.{w', w} uliftFunctor.{v, u} where preservesLimitsOfShape {J} := { preservesLimit := fun {K} => { preserves := fun {c} hc => by @@ -57,93 +57,16 @@ noncomputable instance : CreatesLimitsOfSize.{w, u} uliftFunctor.{v, u} where variable {J : Type*} [Category J] {K : J ⥤ Type u} {c : Cocone K} (hc : IsColimit c) variable {lc : Cocone (K ⋙ uliftFunctor.{v, u})} -/-- Given a subset of the cocone point of a cocone over the lifted functor, - produce a cocone over the original functor. -/ -def coconeOfSet (ls : Set lc.pt) : Cocone K where - pt := ULift Prop - ι := - { app := fun j x ↦ ⟨lc.ι.app j ⟨x⟩ ∈ ls⟩ - naturality := fun i j f ↦ by dsimp only; rw [← lc.w f]; rfl } - -/-- Given a subset of the cocone point of a cocone over the lifted functor, - produce a subset of the cocone point of a colimit cocone over the original functor. -/ -def descSet (ls : Set lc.pt) : Set c.pt := {x | (hc.desc (coconeOfSet ls) x).down} - -/-- Characterization the map `descSet hc`: the image of an element in a vertex of the original - diagram in the cocone point lies in `descSet hc ls` if and only if the image of the corresponding - element in the lifted diagram lie in `ls`. -/ -lemma descSet_spec (s : Set c.pt) (ls : Set lc.pt) : - descSet hc ls = s ↔ ∀ j x, lc.ι.app j ⟨x⟩ ∈ ls ↔ c.ι.app j x ∈ s := by - refine ⟨?_, fun he ↦ funext fun x ↦ ?_⟩ - · rintro rfl j x - exact (congr_arg ULift.down (congr_fun (hc.fac (coconeOfSet ls) j) x).symm).to_iff - · refine (congr_arg ULift.down (congr_fun (hc.uniq (coconeOfSet ls) (⟨· ∈ s⟩) fun j ↦ ?_) x)).symm - ext y; exact congr_arg ULift.up (propext (he j y).symm) - -lemma mem_descSet_singleton {x : lc.pt} {j : J} {y : K.obj j} : - c.ι.app j y ∈ descSet hc {x} ↔ lc.ι.app j ⟨y⟩ = x := - ((descSet_spec hc _ {x}).mp rfl j y).symm - -variable (lc) - -lemma descSet_univ : descSet hc (@Set.univ lc.pt) = Set.univ := by simp [descSet_spec] - -lemma iUnion_descSet_singleton : ⋃ x : lc.pt, descSet hc {x} = Set.univ := by - rw [← descSet_univ hc lc, eq_comm, descSet_spec] - intro j x - erw [true_iff, Set.mem_iUnion] - use lc.ι.app j ⟨x⟩ - rw [mem_descSet_singleton] - -lemma descSet_empty : descSet hc (∅ : Set lc.pt) = ∅ := by simp [descSet_spec] - -lemma descSet_inter_of_ne (x y : lc.pt) (hn : x ≠ y) : descSet hc {x} ∩ descSet hc {y} = ∅ := by - rw [eq_comm, ← descSet_empty hc lc, descSet_spec] - intro j z - erw [false_iff] - rintro ⟨hx, hy⟩ - rw [mem_descSet_singleton] at hx hy - exact hn (hx ▸ hy) - -lemma existsUnique_mem_descSet (x : c.pt) : ∃! y : lc.pt, x ∈ descSet hc {y} := - existsUnique_of_exists_of_unique - (Set.mem_iUnion.mp <| Set.eq_univ_iff_forall.mp (iUnion_descSet_singleton hc lc) x) - fun y₁ y₂ h₁ h₂ ↦ by_contra fun hn ↦ - Set.eq_empty_iff_forall_not_mem.1 (descSet_inter_of_ne hc lc y₁ y₂ hn) x ⟨h₁, h₂⟩ - -@[deprecated (since := "2024-12-17")] alias exists_unique_mem_descSet := existsUnique_mem_descSet - -/-- Given a colimit cocone in `Type u` and an arbitrary cocone over the diagram lifted to - `Type (max u v)`, produce a function from the cocone point of the colimit cocone to the - cocone point of the other cocone, that witnesses the colimit cocone also being a colimit - in the higher universe. -/ -noncomputable def descFun (x : c.pt) : lc.pt := (existsUnique_mem_descSet hc lc x).exists.choose - -lemma descFun_apply_spec {x : c.pt} {y : lc.pt} : descFun hc lc x = y ↔ x ∈ descSet hc {y} := - have hu := existsUnique_mem_descSet hc lc x - have hm := hu.exists.choose_spec - ⟨fun he ↦ he ▸ hm, hu.unique hm⟩ - -lemma descFun_spec (f : c.pt → lc.pt) : - f = descFun hc lc ↔ ∀ j, f ∘ c.ι.app j = lc.ι.app j ∘ ULift.up := by - refine ⟨?_, fun he ↦ funext fun x ↦ ((descFun_apply_spec hc lc).mpr ?_).symm⟩ - · rintro rfl j; ext - apply (descFun_apply_spec hc lc).mpr - rw [mem_descSet_singleton]; rfl - · rw [← (jointly_surjective_of_isColimit hc x).choose_spec.choose_spec, mem_descSet_singleton] - exact (congr_fun (he _) _).symm - /-- The functor `uliftFunctor : Type u ⥤ Type (max u v)` preserves colimits of arbitrary size. -/ noncomputable instance : PreservesColimitsOfSize.{w', w} uliftFunctor.{v, u} where preservesColimitsOfShape {J _} := { preservesColimit := fun {F} ↦ - { preserves := fun {c} hc ↦ ⟨{ - desc := fun lc x ↦ descFun hc lc x.down - fac := fun lc j ↦ by ext ⟨⟩; apply congr_fun ((descFun_spec hc lc _).mp rfl j) - uniq := fun lc f hf ↦ by ext ⟨⟩; apply congr_fun ((descFun_spec hc lc (f ∘ ULift.up)).mpr - fun j ↦ funext fun y ↦ congr_fun (hf j) ⟨y⟩) }⟩ } } + { preserves := fun {c} hc ↦ by + rw [isColimit_iff_bijective_desc, ← Function.Bijective.of_comp_iff _ + (quotQuotUliftEquiv F).bijective, Quot.desc_quotQuotUliftEquiv] + exact ULift.up_bijective.comp ((isColimit_iff_bijective_desc c).mp (Nonempty.intro hc)) } } /-- The functor `uliftFunctor : Type u ⥤ Type (max u v)` creates `u`-small colimits. diff --git a/Mathlib/CategoryTheory/Limits/Types.lean b/Mathlib/CategoryTheory/Limits/Types.lean index 0eff7c3add87e..fd38583178553 100644 --- a/Mathlib/CategoryTheory/Limits/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Types.lean @@ -23,7 +23,7 @@ and the type `lim Hom(F·, X)`. open CategoryTheory CategoryTheory.Limits -universe v u w +universe u' v u w namespace CategoryTheory.Limits @@ -357,6 +357,56 @@ lemma Quot.ι_desc (j : J) (x : F.obj j) : Quot.desc c (Quot.ι F j x) = c.ι.ap lemma Quot.map_ι {j j' : J} {f : j ⟶ j'} (x : F.obj j) : Quot.ι F j' (F.map f x) = Quot.ι F j x := (Quot.sound ⟨f, rfl⟩).symm +/-- +The obvious map from `Quot F` to `Quot (F ⋙ uliftFunctor.{u'})`. +-/ +def quotToQuotUlift (F : J ⥤ Type u) : Quot F → Quot (F ⋙ uliftFunctor.{u'}) := by + refine Quot.lift (fun ⟨j, x⟩ ↦ Quot.ι _ j (ULift.up x)) ?_ + intro ⟨j, x⟩ ⟨j', y⟩ ⟨(f : j ⟶ j'), (eq : y = F.map f x)⟩ + dsimp + have eq : ULift.up y = (F ⋙ uliftFunctor.{u'}).map f (ULift.up x) := by + rw [eq] + dsimp + rw [eq, Quot.map_ι] + +@[simp] +lemma quotToQuotUlift_ι (F : J ⥤ Type u) (j : J) (x : F.obj j) : + quotToQuotUlift F (Quot.ι F j x) = Quot.ι _ j (ULift.up x) := by + dsimp [quotToQuotUlift, Quot.ι] + +/-- +The obvious map from `Quot (F ⋙ uliftFunctor.{u'})` to `Quot F`. +-/ +def quotUliftToQuot (F : J ⥤ Type u) : Quot (F ⋙ uliftFunctor.{u'}) → Quot F := + Quot.lift (fun ⟨j, x⟩ ↦ Quot.ι _ j x.down) + (fun ⟨_, x⟩ ⟨_, y⟩ ⟨f, (eq : y = ULift.up (F.map f x.down))⟩ ↦ by simp [eq, Quot.map_ι]) + +@[simp] +lemma quotUliftToQuot_ι (F : J ⥤ Type u) (j : J) (x : (F ⋙ uliftFunctor.{u'}).obj j) : + quotUliftToQuot F (Quot.ι _ j x) = Quot.ι F j x.down := by + dsimp [quotUliftToQuot, Quot.ι] + +/-- +The equivalence between `Quot F` and `Quot (F ⋙ uliftFunctor.{u'})`. +-/ +@[simp] +def quotQuotUliftEquiv (F : J ⥤ Type u) : Quot F ≃ Quot (F ⋙ uliftFunctor.{u'}) where + toFun := quotToQuotUlift F + invFun := quotUliftToQuot F + left_inv x := by + obtain ⟨j, y, rfl⟩ := Quot.jointly_surjective x + rw [quotToQuotUlift_ι, quotUliftToQuot_ι] + right_inv x := by + obtain ⟨j, y, rfl⟩ := Quot.jointly_surjective x + rw [quotUliftToQuot_ι, quotToQuotUlift_ι] + rfl + +lemma Quot.desc_quotQuotUliftEquiv {F : J ⥤ Type u} (c : Cocone F) : + Quot.desc (uliftFunctor.{u'}.mapCocone c) ∘ quotQuotUliftEquiv F = ULift.up ∘ Quot.desc c := by + ext x + obtain ⟨_, _, rfl⟩ := Quot.jointly_surjective x + dsimp + /-- (implementation detail) A function `Quot F → α` induces a cocone on `F` as long as the universes work out. -/ @[simps] From e58d5d7294c575f1d97908377baa540bcf193787 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Mon, 6 Jan 2025 11:22:24 +0000 Subject: [PATCH 175/189] chore(RingTheory/Finiteness): rename Module.Finite.out (#20506) Right now it's impossible to do ``` -- finite projective modules class IsFiniteProjective extends Module.Projective R M, Module.Finite R M : Prop where -- oops ``` because of a nameclash (both axioms are called `out`). --- Mathlib/Algebra/Module/FinitePresentation.lean | 8 ++++---- Mathlib/Algebra/Module/Torsion.lean | 2 +- Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean | 2 +- Mathlib/RingTheory/Artinian/Module.lean | 2 +- Mathlib/RingTheory/Finiteness/Basic.lean | 4 ++-- Mathlib/RingTheory/Finiteness/Defs.lean | 6 +++--- Mathlib/RingTheory/Finiteness/Nilpotent.lean | 2 +- Mathlib/RingTheory/Finiteness/TensorProduct.lean | 4 ++-- Mathlib/RingTheory/Flat/EquationalCriterion.lean | 2 +- Mathlib/RingTheory/LocalProperties/Projective.lean | 2 +- Mathlib/RingTheory/LocalRing/Module.lean | 4 ++-- Mathlib/RingTheory/Nakayama.lean | 2 +- 12 files changed, 20 insertions(+), 20 deletions(-) diff --git a/Mathlib/Algebra/Module/FinitePresentation.lean b/Mathlib/Algebra/Module/FinitePresentation.lean index 468b2cc116e18..f956657b81585 100644 --- a/Mathlib/Algebra/Module/FinitePresentation.lean +++ b/Mathlib/Algebra/Module/FinitePresentation.lean @@ -187,7 +187,7 @@ lemma Module.FinitePresentation.fg_ker [Module.Finite R M] exact ⟨_, hy, by simp⟩ apply Submodule.fg_of_fg_map_of_fg_inf_ker f.range.mkQ · rw [this] - exact Module.Finite.out + exact Module.Finite.fg_top · rw [Submodule.ker_mkQ, inf_comm, ← Submodule.map_comap_eq, ← LinearMap.ker_comp, hf] exact hs'.map f @@ -201,8 +201,8 @@ lemma Module.finitePresentation_of_ker [Module.FinitePresentation R N] Module.FinitePresentation R M := by obtain ⟨s, hs⟩ : (⊤ : Submodule R M).FG := by apply Submodule.fg_of_fg_map_of_fg_inf_ker l - · rw [Submodule.map_top, LinearMap.range_eq_top.mpr hl]; exact Module.Finite.out - · rw [top_inf_eq, ← Submodule.fg_top]; exact Module.Finite.out + · rw [Submodule.map_top, LinearMap.range_eq_top.mpr hl]; exact Module.Finite.fg_top + · rw [top_inf_eq, ← Submodule.fg_top]; exact Module.Finite.fg_top refine ⟨s, hs, ?_⟩ let π := Finsupp.linearCombination R ((↑) : s → M) have H : Function.Surjective π := @@ -250,7 +250,7 @@ instance {A} [CommRing A] [Algebra R A] [Module.FinitePresentation R M] : apply Submodule.FG.map have : Module.Finite R (LinearMap.ker f) := ⟨(Submodule.fg_top _).mpr (Module.FinitePresentation.fg_ker f hf)⟩ - exact Module.Finite.out (R := A) (M := A ⊗[R] LinearMap.ker f) + exact Module.Finite.fg_top (R := A) (M := A ⊗[R] LinearMap.ker f) open TensorProduct in lemma FinitePresentation.of_isBaseChange diff --git a/Mathlib/Algebra/Module/Torsion.lean b/Mathlib/Algebra/Module/Torsion.lean index 08b33bcef771a..c7ce5c683b548 100644 --- a/Mathlib/Algebra/Module/Torsion.lean +++ b/Mathlib/Algebra/Module/Torsion.lean @@ -691,7 +691,7 @@ variable {R M} theorem _root_.Submodule.annihilator_top_inter_nonZeroDivisors [Module.Finite R M] (hM : Module.IsTorsion R M) : ((⊤ : Submodule R M).annihilator : Set R) ∩ R⁰ ≠ ∅ := by - obtain ⟨S, hS⟩ := ‹Module.Finite R M›.out + obtain ⟨S, hS⟩ := ‹Module.Finite R M›.fg_top refine Set.Nonempty.ne_empty ⟨_, ?_, (∏ x ∈ S, (@hM x).choose : R⁰).prop⟩ rw [Submonoid.coe_finset_prod, SetLike.mem_coe, ← hS, mem_annihilator_span] intro n diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean index a6138962526ff..db92f36bbab1d 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean @@ -210,7 +210,7 @@ theorem LinearMap.exists_monic_and_coeff_mem_pow_and_aeval_eq_zero_of_range_le_s cases subsingleton_or_nontrivial R · exact ⟨0, Polynomial.monic_of_subsingleton _, by simp⟩ obtain ⟨s : Finset M, hs : Submodule.span R (s : Set M) = ⊤⟩ := - Module.Finite.out (R := R) (M := M) + Module.Finite.fg_top (R := R) (M := M) -- Porting note: `H` was `rfl` obtain ⟨A, H, h⟩ := Matrix.isRepresentation.toEnd_exists_mem_ideal R ((↑) : s → M) diff --git a/Mathlib/RingTheory/Artinian/Module.lean b/Mathlib/RingTheory/Artinian/Module.lean index 53f91c83eee98..61c8771b69611 100644 --- a/Mathlib/RingTheory/Artinian/Module.lean +++ b/Mathlib/RingTheory/Artinian/Module.lean @@ -390,7 +390,7 @@ theorem isArtinian_of_fg_of_artinian {R M} [Ring R] [AddCommGroup M] [Module R M instance isArtinian_of_fg_of_artinian' {R M} [Ring R] [AddCommGroup M] [Module R M] [IsArtinianRing R] [Module.Finite R M] : IsArtinian R M := - have : IsArtinian R (⊤ : Submodule R M) := isArtinian_of_fg_of_artinian _ Module.Finite.out + have : IsArtinian R (⊤ : Submodule R M) := isArtinian_of_fg_of_artinian _ Module.Finite.fg_top isArtinian_of_linearEquiv (LinearEquiv.ofTop (⊤ : Submodule R M) rfl) theorem IsArtinianRing.of_finite (R S) [CommRing R] [Ring S] [Algebra R S] diff --git a/Mathlib/RingTheory/Finiteness/Basic.lean b/Mathlib/RingTheory/Finiteness/Basic.lean index 73248997ca909..bb1058f07fbc1 100644 --- a/Mathlib/RingTheory/Finiteness/Basic.lean +++ b/Mathlib/RingTheory/Finiteness/Basic.lean @@ -251,13 +251,13 @@ theorem equiv_iff (e : M ≃ₗ[R] N) : Module.Finite R M ↔ Module.Finite R N instance ulift [Module.Finite R M] : Module.Finite R (ULift M) := equiv ULift.moduleEquiv.symm -theorem iff_fg {N : Submodule R M} : Module.Finite R N ↔ N.FG := Module.finite_def.trans (fg_top _) +theorem iff_fg {N : Submodule R M} : Module.Finite R N ↔ N.FG := Module.finite_def.trans N.fg_top variable (R M) instance bot : Module.Finite R (⊥ : Submodule R M) := iff_fg.mpr fg_bot -instance top [Module.Finite R M] : Module.Finite R (⊤ : Submodule R M) := iff_fg.mpr out +instance top [Module.Finite R M] : Module.Finite R (⊤ : Submodule R M) := iff_fg.mpr fg_top variable {M} diff --git a/Mathlib/RingTheory/Finiteness/Defs.lean b/Mathlib/RingTheory/Finiteness/Defs.lean index fd8c45bb0be8d..9c52d2ce2bd80 100644 --- a/Mathlib/RingTheory/Finiteness/Defs.lean +++ b/Mathlib/RingTheory/Finiteness/Defs.lean @@ -103,9 +103,9 @@ variable (R A B M N : Type*) /-- A module over a semiring is `Module.Finite` if it is finitely generated as a module. -/ protected class Module.Finite [Semiring R] [AddCommMonoid M] [Module R M] : Prop where - out : (⊤ : Submodule R M).FG + fg_top : (⊤ : Submodule R M).FG -attribute [inherit_doc Module.Finite] Module.Finite.out +attribute [inherit_doc Module.Finite] Module.Finite.fg_top namespace Module @@ -131,7 +131,7 @@ variable {R M N} /-- See also `Module.Finite.exists_fin'`. -/ lemma exists_fin [Module.Finite R M] : ∃ (n : ℕ) (s : Fin n → M), Submodule.span R (range s) = ⊤ := - Submodule.fg_iff_exists_fin_generating_family.mp out + Submodule.fg_iff_exists_fin_generating_family.mp fg_top end Finite diff --git a/Mathlib/RingTheory/Finiteness/Nilpotent.lean b/Mathlib/RingTheory/Finiteness/Nilpotent.lean index 17b91ef86e77f..b2e0915fb0405 100644 --- a/Mathlib/RingTheory/Finiteness/Nilpotent.lean +++ b/Mathlib/RingTheory/Finiteness/Nilpotent.lean @@ -20,7 +20,7 @@ namespace Finite theorem Module.End.isNilpotent_iff_of_finite [Module.Finite R M] {f : End R M} : IsNilpotent f ↔ ∀ m : M, ∃ n : ℕ, (f ^ n) m = 0 := by refine ⟨fun ⟨n, hn⟩ m ↦ ⟨n, by simp [hn]⟩, fun h ↦ ?_⟩ - rcases Module.Finite.out (R := R) (M := M) with ⟨S, hS⟩ + rcases Module.Finite.fg_top (R := R) (M := M) with ⟨S, hS⟩ choose g hg using h use Finset.sup S g ext m diff --git a/Mathlib/RingTheory/Finiteness/TensorProduct.lean b/Mathlib/RingTheory/Finiteness/TensorProduct.lean index e1a686b2c3209..abfabac2a55e3 100644 --- a/Mathlib/RingTheory/Finiteness/TensorProduct.lean +++ b/Mathlib/RingTheory/Finiteness/TensorProduct.lean @@ -85,7 +85,7 @@ noncomputable local instance instance Module.Finite.base_change [CommSemiring R] [Semiring A] [Algebra R A] [AddCommMonoid M] [Module R M] [h : Module.Finite R M] : Module.Finite A (TensorProduct R A M) := by classical - obtain ⟨s, hs⟩ := h.out + obtain ⟨s, hs⟩ := h.fg_top refine ⟨⟨s.image (TensorProduct.mk R A M 1), eq_top_iff.mpr ?_⟩⟩ rintro x - induction x with @@ -103,7 +103,7 @@ instance Module.Finite.base_change [CommSemiring R] [Semiring A] [Algebra R A] [ instance Module.Finite.tensorProduct [CommSemiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid N] [Module R N] [hM : Module.Finite R M] [hN : Module.Finite R N] : Module.Finite R (TensorProduct R M N) where - out := (TensorProduct.map₂_mk_top_top_eq_top R M N).subst (hM.out.map₂ _ hN.out) + fg_top := (TensorProduct.map₂_mk_top_top_eq_top R M N).subst (hM.fg_top.map₂ _ hN.fg_top) end ModuleAndAlgebra diff --git a/Mathlib/RingTheory/Flat/EquationalCriterion.lean b/Mathlib/RingTheory/Flat/EquationalCriterion.lean index 763d9dc2c83db..df38f6aa41c23 100644 --- a/Mathlib/RingTheory/Flat/EquationalCriterion.lean +++ b/Mathlib/RingTheory/Flat/EquationalCriterion.lean @@ -273,7 +273,7 @@ theorem exists_factorization_of_comp_eq_zero_of_free [Flat R M] {K N : Type u} [ use κ₂, hκ₂, a₂ ∘ₗ a₁, y₂ simp_rw [comp_assoc] exact ⟨trivial, sup_le (ha₁.trans (ker_le_ker_comp _ _)) ha₂⟩ - convert this ⊤ Finite.out + convert this ⊤ Finite.fg_top simp only [top_le_iff, ker_eq_top] /-- Every homomorphism from a finitely presented module to a flat module factors through a finite diff --git a/Mathlib/RingTheory/LocalProperties/Projective.lean b/Mathlib/RingTheory/LocalProperties/Projective.lean index bf3d45d29deed..dc9a8726d2715 100644 --- a/Mathlib/RingTheory/LocalProperties/Projective.lean +++ b/Mathlib/RingTheory/LocalProperties/Projective.lean @@ -121,7 +121,7 @@ theorem Module.projective_of_localization_maximal (H : ∀ (I : Ideal R) (_ : I. Module.Projective (Localization.AtPrime I) (LocalizedModule I.primeCompl M)) [Module.FinitePresentation R M] : Module.Projective R M := by have : Module.Finite R M := by infer_instance - have : (⊤ : Submodule R M).FG := this.out + have : (⊤ : Submodule R M).FG := this.fg_top have : ∃ (s : Finset M), _ := this obtain ⟨s, hs⟩ := this let N := s →₀ R diff --git a/Mathlib/RingTheory/LocalRing/Module.lean b/Mathlib/RingTheory/LocalRing/Module.lean index 8aec28a9bca3d..fdc9bee0bd9c1 100644 --- a/Mathlib/RingTheory/LocalRing/Module.lean +++ b/Mathlib/RingTheory/LocalRing/Module.lean @@ -61,7 +61,7 @@ theorem map_mkQ_eq {N₁ N₂ : Submodule R M} (h : N₁ ≤ N₂) (h' : N₂.FG theorem map_mkQ_eq_top {N : Submodule R M} [Module.Finite R M] : N.map (Submodule.mkQ (𝔪 • ⊤)) = ⊤ ↔ N = ⊤ := by - rw [← map_mkQ_eq (N₁ := N) le_top Module.Finite.out, Submodule.map_top, Submodule.range_mkQ] + rw [← map_mkQ_eq (N₁ := N) le_top Module.Finite.fg_top, Submodule.map_top, Submodule.range_mkQ] theorem map_tensorProduct_mk_eq_top {N : Submodule R M} [Module.Finite R M] : N.map (TensorProduct.mk R k M 1) = ⊤ ↔ N = ⊤ := by @@ -267,7 +267,7 @@ theorem free_of_lTensor_residueField_injective (hg : Surjective g) (h : Exact f (hf : Function.Injective (f.lTensor k)) : Module.Free R P := by have := Module.finitePresentation_of_free_of_surjective g hg - (by rw [h.linearMap_ker_eq, LinearMap.range_eq_map]; exact (Module.Finite.out).map f) + (by rw [h.linearMap_ker_eq, LinearMap.range_eq_map]; exact (Module.Finite.fg_top).map f) apply free_of_maximalIdeal_rTensor_injective rw [← LinearMap.lTensor_inj_iff_rTensor_inj] apply lTensor_injective_of_exact_of_exact_of_rTensor_injective diff --git a/Mathlib/RingTheory/Nakayama.lean b/Mathlib/RingTheory/Nakayama.lean index 0227b2a554422..ad4d01141b8f1 100644 --- a/Mathlib/RingTheory/Nakayama.lean +++ b/Mathlib/RingTheory/Nakayama.lean @@ -84,7 +84,7 @@ lemma eq_bot_of_set_smul_eq_of_subset_jacobson_annihilator {s : Set R} lemma top_ne_ideal_smul_of_le_jacobson_annihilator [Nontrivial M] [Module.Finite R M] {I} (h : I ≤ (Module.annihilator R M).jacobson) : (⊤ : Submodule R M) ≠ I • ⊤ := fun H => top_ne_bot <| - eq_bot_of_eq_ideal_smul_of_le_jacobson_annihilator Module.Finite.out H <| + eq_bot_of_eq_ideal_smul_of_le_jacobson_annihilator Module.Finite.fg_top H <| (congrArg (I ≤ Ideal.jacobson ·) annihilator_top).mpr h open Pointwise in From ae4ed0224c0525a639077956f6520becf0114f7e Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Mon, 6 Jan 2025 11:55:12 +0000 Subject: [PATCH 176/189] feat(Topology/Algebra/InfiniteSum/NatInt): Add pnat lems (#16544) Add some basic results about tsums indexed by pnat. Co-authored-by: Chris Birkbeck --- Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean b/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean index d1448d49fe4b3..662314f0bb691 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean @@ -514,3 +514,16 @@ lemma multipliable_int_iff_multipliable_nat_and_neg {f : ℤ → G} : end UniformGroup end Int + +section pnat + +@[to_additive] +theorem pnat_multipliable_iff_multipliable_succ {α : Type*} [TopologicalSpace α] [CommMonoid α] + {f : ℕ → α} : Multipliable (fun x : ℕ+ => f x) ↔ Multipliable fun x : ℕ => f (x + 1) := + Equiv.pnatEquivNat.symm.multipliable_iff.symm + +@[to_additive] +theorem tprod_pnat_eq_tprod_succ {α : Type*} [TopologicalSpace α] [CommMonoid α] (f : ℕ → α) : + ∏' n : ℕ+, f n = ∏' n, f (n + 1) := (Equiv.pnatEquivNat.symm.tprod_eq _).symm + +end pnat From 5b6f2c614086f75cce987fd1f75deefbae4b43bf Mon Sep 17 00:00:00 2001 From: smorel394 Date: Mon, 6 Jan 2025 12:15:39 +0000 Subject: [PATCH 177/189] refactor(Algebra/Category/Grp/Colimits): change the construction of colimits in `AddCommGrp` (#20474) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The old version of this file has a remark saying "TODO: In fact, in `AddCommGrp` there is a much nicer model of colimits as quotients of finitely supported functions, and we really should implement this as well (or instead)." This PR does precisely this (choosing the "instead" option). In fact, given a functor `F : J ⥤ AddCommGrp.{w}`, it constructs a candidate for the colimit of `F` as a quotient of `DFinsupp (fun j ↦ F.obj j)`, without any smallness assumptions on `J`. It then proves that this quotient is a colimit of `F` if it is `w`-small, then deduces that `AddCommGrp.{w}` has all small colimits. Co-authored-by: smorel394 <67864981+smorel394@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Algebra/Category/Grp/Colimits.lean | 352 ++++++++---------- Mathlib/Algebra/Homology/LocalCohomology.lean | 4 +- Mathlib/Data/DFinsupp/Small.lean | 28 ++ 4 files changed, 176 insertions(+), 209 deletions(-) create mode 100644 Mathlib/Data/DFinsupp/Small.lean diff --git a/Mathlib.lean b/Mathlib.lean index f692eadfdec90..99606d29a64b3 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2456,6 +2456,7 @@ import Mathlib.Data.DFinsupp.NeLocus import Mathlib.Data.DFinsupp.Notation import Mathlib.Data.DFinsupp.Order import Mathlib.Data.DFinsupp.Sigma +import Mathlib.Data.DFinsupp.Small import Mathlib.Data.DFinsupp.Submonoid import Mathlib.Data.DFinsupp.WellFounded import Mathlib.Data.DList.Instances diff --git a/Mathlib/Algebra/Category/Grp/Colimits.lean b/Mathlib/Algebra/Category/Grp/Colimits.lean index 1164b843e4f0b..a7d7824126397 100644 --- a/Mathlib/Algebra/Category/Grp/Colimits.lean +++ b/Mathlib/Algebra/Category/Grp/Colimits.lean @@ -1,37 +1,29 @@ /- Copyright (c) 2019 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kim Morrison +Authors: Kim Morrison, Sophie Morel -/ import Mathlib.Algebra.Category.Grp.Preadditive -import Mathlib.CategoryTheory.Limits.Shapes.Kernels -import Mathlib.CategoryTheory.Limits.Shapes.FiniteLimits +import Mathlib.Algebra.Equiv.TransferInstance import Mathlib.CategoryTheory.ConcreteCategory.Elementwise +import Mathlib.Data.DFinsupp.BigOperators +import Mathlib.Data.DFinsupp.Small import Mathlib.GroupTheory.QuotientGroup.Defs - /-! # The category of additive commutative groups has all colimits. -This file uses a "pre-automated" approach, just as for `Algebra.Category.MonCat.Colimits`. -It is a very uniform approach, that conceivably could be synthesised directly -by a tactic that analyses the shape of `AddCommGroup` and `MonoidHom`. +This file constructs colimits in the categpry of additive commutative groups, as +quotients of finitely supported functions. -TODO: -In fact, in `AddCommGrp` there is a much nicer model of colimits as quotients -of finitely supported functions, and we really should implement this as well (or instead). -/ - universe w u v open CategoryTheory Limits --- [ROBOT VOICE]: --- You should pretend for now that this file was automatically generated. --- It follows the same template as colimits in Mon. namespace AddCommGrp -variable {J : Type u} [Category.{v} J] (F : J ⥤ AddCommGrp.{max u v w}) +variable {J : Type u} [Category.{v} J] (F : J ⥤ AddCommGrp.{w}) namespace Colimits @@ -42,220 +34,168 @@ then taking the quotient by the abelian group laws within each abelian group, and the identifications given by the morphisms in the diagram. -/ -/-- An inductive type representing all group expressions (without relations) -on a collection of types indexed by the objects of `J`. --/ -inductive Prequotient - -- There's always `of` - | of : ∀ (j : J) (_ : F.obj j), Prequotient - -- Then one generator for each operation - | zero : Prequotient - | neg : Prequotient → Prequotient - | add : Prequotient → Prequotient → Prequotient - -instance : Inhabited (Prequotient.{w} F) := - ⟨Prequotient.zero⟩ - -open Prequotient - -/-- The relation on `Prequotient` saying when two expressions are equal -because of the abelian group laws, or -because one element is mapped to another by a morphism in the diagram. +/-- +The relations between elements of the direct sum of the `F.obj j` given by the +morphisms in the diagram `J`. -/ -inductive Relation : Prequotient.{w} F → Prequotient.{w} F → Prop - -- Make it an equivalence relation: - | refl : ∀ x, Relation x x - | symm : ∀ (x y) (_ : Relation x y), Relation y x - | trans : ∀ (x y z) (_ : Relation x y) (_ : Relation y z), Relation x z - -- There's always a `map` relation - | map : ∀ (j j' : J) (f : j ⟶ j') (x : F.obj j), Relation (Prequotient.of j' (F.map f x)) - (Prequotient.of j x) - -- Then one relation per operation, describing the interaction with `of` - | zero : ∀ j, Relation (Prequotient.of j 0) zero - | neg : ∀ (j) (x : F.obj j), Relation (Prequotient.of j (-x)) (neg (Prequotient.of j x)) - | add : ∀ (j) (x y : F.obj j), Relation (Prequotient.of j (x + y)) (add (Prequotient.of j x) - (Prequotient.of j y)) - -- Then one relation per argument of each operation - | neg_1 : ∀ (x x') (_ : Relation x x'), Relation (neg x) (neg x') - | add_1 : ∀ (x x' y) (_ : Relation x x'), Relation (add x y) (add x' y) - | add_2 : ∀ (x y y') (_ : Relation y y'), Relation (add x y) (add x y') - -- And one relation per axiom - | zero_add : ∀ x, Relation (add zero x) x - | add_zero : ∀ x, Relation (add x zero) x - | 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)) +abbrev Relations [DecidableEq J] : AddSubgroup (DFinsupp (fun j ↦ F.obj j)) := + AddSubgroup.closure {x | ∃ (j j' : J) (u : j ⟶ j') (a : F.obj j), + x = DFinsupp.single j' (F.map u a) - DFinsupp.single j a} /-- -The setoid corresponding to group expressions modulo abelian group relations and identifications. +The candidate for the colimit of `F`, defined as the quotient of the direct sum +of the commutative groups `F.obj j` by the relations given by the morphisms in +the diagram. -/ -def colimitSetoid : Setoid (Prequotient.{w} F) where - r := Relation F - iseqv := ⟨Relation.refl, fun r => Relation.symm _ _ r, fun r => Relation.trans _ _ _ r⟩ +def Quot [DecidableEq J] : Type (max u w) := + DFinsupp (fun j ↦ F.obj j) ⧸ Relations F -attribute [instance] colimitSetoid +instance [DecidableEq J] : AddCommGroup (Quot F) := + QuotientAddGroup.Quotient.addCommGroup (Relations F) -/-- The underlying type of the colimit of a diagram in `AddCommGrp`. +/-- Inclusion of `F.obj j` into the candidate colimit. -/ -def ColimitType : Type max u v w := - Quotient (colimitSetoid.{w} F) - -instance : Zero (ColimitType.{w} F) where - zero := Quotient.mk _ zero - -instance : Neg (ColimitType.{w} F) where - neg := Quotient.map neg Relation.neg_1 - -instance : Add (ColimitType.{w} F) where - add := Quotient.map₂ add <| fun _x x' rx y _y' ry => - Setoid.trans (Relation.add_1 _ _ y rx) (Relation.add_2 x' _ _ ry) - -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 _ - 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 _ _ _ - nsmul := nsmulRec - zsmul := zsmulRec - -instance ColimitTypeInhabited : Inhabited (ColimitType.{w} F) := ⟨0⟩ - -@[simp] -theorem quot_zero : Quot.mk Setoid.r zero = (0 : ColimitType.{w} F) := - rfl +def Quot.ι [DecidableEq J] (j : J) : F.obj j →+ Quot F := + (QuotientAddGroup.mk' _).comp (DFinsupp.singleAddHom (fun j ↦ F.obj j) j) + +lemma Quot.addMonoidHom_ext [DecidableEq J] {α : Type*} [AddMonoid α] {f g : Quot F →+ α} + (h : ∀ (j : J) (x : F.obj j), f (Quot.ι F j x) = g (Quot.ι F j x)) : f = g := + QuotientAddGroup.addMonoidHom_ext _ (DFinsupp.addHom_ext h) + +variable (c : Cocone F) + +/-- (implementation detail) Part of the universal property of the colimit cocone, but without + assuming that `Quot F` lives in the correct universe. -/ +def Quot.desc [DecidableEq J] : Quot.{w} F →+ c.pt := by + refine QuotientAddGroup.lift _ (DFinsupp.sumAddHom c.ι.app) ?_ + dsimp + rw [AddSubgroup.closure_le] + intro _ ⟨_, _, _, _, eq⟩ + rw [eq] + simp only [SetLike.mem_coe, AddMonoidHom.mem_ker, map_sub, DFinsupp.sumAddHom_single] + change (F.map _ ≫ c.ι.app _) _ - _ = 0 + rw [c.ι.naturality] + simp only [Functor.const_obj_obj, Functor.const_obj_map, Category.comp_id, sub_self] @[simp] -theorem quot_neg (x) : - -- Porting note: force Lean to treat `ColimitType F` no as `Quot _` - (by exact Quot.mk Setoid.r (neg x) : ColimitType.{w} F) = - -(by exact Quot.mk Setoid.r x) := - rfl +lemma Quot.ι_desc [DecidableEq J] (j : J) (x : F.obj j) : + Quot.desc F c (Quot.ι F j x) = c.ι.app j x := by + dsimp [desc, ι] + erw [QuotientAddGroup.lift_mk'] + simp @[simp] -theorem quot_add (x y) : - (by exact Quot.mk Setoid.r (add x y) : ColimitType.{w} F) = - -- Porting note: force Lean to treat `ColimitType F` no as `Quot _` - (by exact Quot.mk Setoid.r x) + (by exact Quot.mk Setoid.r y) := - rfl - -/-- The bundled abelian group giving the colimit of a diagram. -/ -def colimit : AddCommGrp := - AddCommGrp.of (ColimitType.{w} F) - -/-- The function from a given abelian group in the diagram to the colimit abelian group. -/ -def coconeFun (j : J) (x : F.obj j) : ColimitType.{w} F := - Quot.mk _ (Prequotient.of j x) - -/-- The group homomorphism from a given abelian group in the diagram to the colimit abelian -group. -/ -def coconeMorphism (j : J) : F.obj j ⟶ colimit.{w} F where - toFun := coconeFun F j - map_zero' := by apply Quot.sound; apply Relation.zero - map_add' := by intros; apply Quot.sound; apply Relation.add +lemma Quot.map_ι [DecidableEq J] {j j' : J} {f : j ⟶ j'} (x : F.obj j) : + Quot.ι F j' (F.map f x) = Quot.ι F j x := by + dsimp [ι] + refine eq_of_sub_eq_zero ?_ + erw [← (QuotientAddGroup.mk' (Relations F)).map_sub, ← AddMonoidHom.mem_ker] + rw [QuotientAddGroup.ker_mk'] + simp only [DFinsupp.singleAddHom_apply] + exact AddSubgroup.subset_closure ⟨j, j', f, x, rfl⟩ + +/-- (implementation detail) A morphism of commutative additive groups `Quot F →+ A` +induces a cocone on `F` as long as the universes work out. +-/ +@[simps] +def toCocone [DecidableEq J] {A : Type w} [AddCommGroup A] (f : Quot F →+ A) : Cocone F where + pt := AddCommGrp.of A + ι := { app := fun j => f.comp (Quot.ι F j) } + +lemma Quot.desc_toCocone_desc [DecidableEq J] {A : Type w} [AddCommGroup A] (f : Quot F →+ A) + (hc : IsColimit c) : (hc.desc (toCocone F f)).comp (Quot.desc F c) = f := by + refine Quot.addMonoidHom_ext F (fun j x ↦ ?_) + rw [AddMonoidHom.comp_apply, ι_desc] + change (c.ι.app j ≫ hc.desc (toCocone F f)) _ = _ + rw [hc.fac] + simp + +lemma Quot.desc_toCocone_desc_app [DecidableEq J] {A : Type w} [AddCommGroup A] (f : Quot F →+ A) + (hc : IsColimit c) (x : Quot F) : hc.desc (toCocone F f) (Quot.desc F c x) = f x := by + conv_rhs => rw [← Quot.desc_toCocone_desc F c f hc] + dsimp -@[simp] -theorem cocone_naturality {j j' : J} (f : j ⟶ j') : - F.map f ≫ coconeMorphism.{w} F j' = coconeMorphism F j := by - ext - apply Quot.sound - apply Relation.map +/-- +If `c` is a cocone of `F` such that `Quot.desc F c` is bijective, then `c` is a colimit +cocone of `F`. +-/ +noncomputable def isColimit_of_bijective_desc [DecidableEq J] + (h : Function.Bijective (Quot.desc F c)) : IsColimit c where + desc s := AddCommGrp.ofHom ((Quot.desc F s).comp (AddEquiv.ofBijective + (Quot.desc F c) h).symm.toAddMonoidHom) + fac s j := by + ext x + dsimp + conv_lhs => erw [← Quot.ι_desc F c j x] + rw [← AddEquiv.ofBijective_apply _ h, AddEquiv.symm_apply_apply] + simp only [Quot.ι_desc, Functor.const_obj_obj] + uniq s m hm := by + ext x + obtain ⟨x, rfl⟩ := h.2 x + dsimp + rw [← AddEquiv.ofBijective_apply _ h, AddEquiv.symm_apply_apply] + suffices eq : m.comp (AddEquiv.ofBijective (Quot.desc F c) h) = Quot.desc F s by + rw [← eq]; rfl + exact Quot.addMonoidHom_ext F (by simp [← hm]) + +/-- (internal implementation) The colimit cocone of a functor `F`, implemented as a quotient of +`DFinsupp (fun j ↦ F.obj j)`, under the assumption that said quotient is small. +-/ +@[simps] +noncomputable def colimitCocone [DecidableEq J] [Small.{w} (Quot.{w} F)] : Cocone F where + pt := AddCommGrp.of (Shrink (Quot F)) + ι := + { app j := + AddCommGrp.ofHom (Shrink.addEquiv.symm.toAddMonoidHom.comp (Quot.ι F j)) + naturality _ _ _ := by + ext + dsimp + simp only [Category.comp_id, ofHom_apply, AddMonoidHom.coe_comp, AddMonoidHom.coe_coe, + Function.comp_apply] + change Shrink.addEquiv.symm _ = _ + rw [Quot.map_ι] } @[simp] -theorem cocone_naturality_components (j j' : J) (f : j ⟶ j') (x : F.obj j) : - (coconeMorphism.{w} F j') (F.map f x) = (coconeMorphism F j) x := by - rw [← cocone_naturality F f] - rfl - -/-- The cocone over the proposed colimit abelian group. -/ -def colimitCocone : Cocone F where - pt := colimit.{w} F - ι := { app := coconeMorphism F } - -/-- The function from the free abelian group on the diagram to the cone point of any other -cocone. -/ -@[simp] -def descFunLift (s : Cocone F) : Prequotient.{w} F → s.pt - | Prequotient.of j x => (s.ι.app j) x - | zero => 0 - | neg x => -descFunLift s x - | add x y => descFunLift s x + descFunLift s y - -/-- The function from the colimit abelian group to the cone point of any other cocone. -/ -def descFun (s : Cocone F) : ColimitType.{w} F → s.pt := by - fapply Quot.lift - · exact descFunLift F s - · intro x y r - induction r with - | refl => rfl - | symm _ _ _ r_ih => exact r_ih.symm - | trans _ _ _ _ _ r_ih_h r_ih_k => exact Eq.trans r_ih_h r_ih_k - | map j j' f x => simpa only [descFunLift, Functor.const_obj_obj] using - DFunLike.congr_fun (s.ι.naturality f) x - | zero => simp - | neg => simp - | add => simp - | neg_1 _ _ _ r_ih => dsimp; rw [r_ih] - | add_1 _ _ _ _ r_ih => dsimp; rw [r_ih] - | add_2 _ _ _ _ r_ih => dsimp; rw [r_ih] - | zero_add => dsimp; rw [zero_add] - | add_zero => dsimp; rw [add_zero] - | neg_add_cancel => dsimp; rw [neg_add_cancel] - | add_comm => dsimp; rw [add_comm] - | add_assoc => dsimp; rw [add_assoc] - -/-- The group homomorphism from the colimit abelian group to the cone point of any other cocone. -/ -def descMorphism (s : Cocone F) : colimit.{w} F ⟶ s.pt where - toFun := descFun F s - map_zero' := rfl - map_add' x y := Quot.induction_on₂ x y fun _ _ ↦ rfl - -/-- Evidence that the proposed colimit is the colimit. -/ -def colimitCoconeIsColimit : IsColimit (colimitCocone.{w} F) where - desc s := descMorphism F s - uniq s m w := DFunLike.ext _ _ fun x => Quot.inductionOn x fun x => by - change (m : ColimitType F →+ s.pt) _ = (descMorphism F s : ColimitType F →+ s.pt) _ - induction x using Prequotient.recOn with - | of j x => exact DFunLike.congr_fun (w j) x - | zero => - dsimp only [quot_zero] - rw [map_zero, map_zero] - | neg x ih => - dsimp only [quot_neg] - rw [map_neg, map_neg, ih] - | add x y ihx ihy => - simp only [quot_add] - rw [m.map_add, (descMorphism F s).map_add, ihx, ihy] +theorem Quot.desc_colimitCocone [DecidableEq J] (F : J ⥤ AddCommGrp.{w}) [Small.{w} (Quot F)] : + Quot.desc F (colimitCocone F) = (Shrink.addEquiv (α := Quot F)).symm.toAddMonoidHom := by + refine Quot.addMonoidHom_ext F (fun j x ↦ ?_) + simp only [colimitCocone_pt, coe_of, AddEquiv.toAddMonoidHom_eq_coe, AddMonoidHom.coe_coe] + erw [Quot.ι_desc] + simp + +/-- (internal implementation) The fact that the candidate colimit cocone constructed in +`colimitCocone` is the colimit. +-/ +noncomputable def colimitCoconeIsColimit [DecidableEq J] [Small.{w} (Quot F)] : + IsColimit (colimitCocone F) := by + refine isColimit_of_bijective_desc F _ ?_ + rw [Quot.desc_colimitCocone] + exact Shrink.addEquiv.symm.bijective end Colimits -lemma hasColimit : HasColimit F := ⟨_, Colimits.colimitCoconeIsColimit.{w} F⟩ - -variable (J) - -lemma hasColimitsOfShape : HasColimitsOfShape J AddCommGrp.{max u v w} where - has_colimit F := hasColimit.{w} F +open Colimits -lemma hasColimitsOfSize : HasColimitsOfSize.{v, u} AddCommGrp.{max u v w} := - ⟨fun _ => hasColimitsOfShape.{w} _⟩ +lemma hasColimit_of_small_quot [DecidableEq J] (h : Small.{w} (Quot F)) : HasColimit F := + ⟨_, colimitCoconeIsColimit F⟩ -instance hasColimits : HasColimits AddCommGrp.{w} := hasColimitsOfSize.{w} +instance [DecidableEq J] [Small.{w} J] : Small.{w} (Quot F) := + small_of_surjective (QuotientAddGroup.mk'_surjective _) -instance : HasColimitsOfSize.{v, v} (AddCommGrpMax.{u, v}) := hasColimitsOfSize.{u} -instance : HasColimitsOfSize.{u, u} (AddCommGrpMax.{u, v}) := hasColimitsOfSize.{v} -instance : HasColimitsOfSize.{u, v} (AddCommGrpMax.{u, v}) := hasColimitsOfSize.{u} -instance : HasColimitsOfSize.{v, u} (AddCommGrpMax.{u, v}) := hasColimitsOfSize.{u} -instance : HasColimitsOfSize.{0, 0} (AddCommGrp.{u}) := hasColimitsOfSize.{u, 0, 0} +instance hasColimit [Small.{w} J] (F : J ⥤ AddCommGrp.{w}) : HasColimit F := by + classical + exact hasColimit_of_small_quot F inferInstance -example : HasColimits AddCommGrpMax.{v, u} := - inferInstance -example : HasColimits AddCommGrpMax.{u, v} := - inferInstance +/-- +If `J` is `w`-small, then any functor `J ⥤ AddCommGrp.{w}` has a colimit. +-/ +instance hasColimitsOfShape [Small.{w} J] : HasColimitsOfShape J (AddCommGrp.{w}) where -example : HasColimits AddCommGrp.{u} := - inferInstance +/-- The category of additive commutative groups has all small colimits. +-/ +instance (priority := 1300) hasColimitsOfSize [UnivLE.{u, w}] : + HasColimitsOfSize.{v, u} (AddCommGrp.{w}) where end AddCommGrp diff --git a/Mathlib/Algebra/Homology/LocalCohomology.lean b/Mathlib/Algebra/Homology/LocalCohomology.lean index ebc2b550a5d65..7b53d4b469f1a 100644 --- a/Mathlib/Algebra/Homology/LocalCohomology.lean +++ b/Mathlib/Algebra/Homology/LocalCohomology.lean @@ -87,9 +87,7 @@ section variable {R : Type max u v} [CommRing R] {D : Type v} [SmallCategory D] lemma hasColimitDiagram (I : D ⥤ Ideal R) (i : ℕ) : - HasColimit (diagram I i) := by - have : HasColimitsOfShape Dᵒᵖ (AddCommGrpMax.{u, v}) := inferInstance - infer_instance + HasColimit (diagram I i) := inferInstance /- In this definition we do not assume any special property of the diagram `I`, but the relevant case diff --git a/Mathlib/Data/DFinsupp/Small.lean b/Mathlib/Data/DFinsupp/Small.lean new file mode 100644 index 0000000000000..04278d86c6531 --- /dev/null +++ b/Mathlib/Data/DFinsupp/Small.lean @@ -0,0 +1,28 @@ +/- +Copyright (c) 2025 Sophie. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sophie Morel +-/ +import Mathlib.Data.DFinsupp.Defs +import Mathlib.Logic.Small.Basic + +/-! +# Smallness of the `DFinsupp` type + +Let `π : ι → Type v`. If `ι` and all the `π i` are `w`-small, this provides a `Small.{w}` +instance on `DFinsupp π`. + +-/ + +universe u v w + +variable {ι : Type u} {π : ι → Type v} [∀ i, Zero (π i)] + +section Small + +instance DFinsupp.small [Small.{w} ι] [∀ (i : ι), Small.{w} (π i)] : + Small.{w} (DFinsupp π) := small_of_injective (f := fun x j ↦ x j) (fun f f' eq ↦ by + ext j + exact congr_fun eq j) + +end Small From baef601f4b4baa6f9682b5f61bda2cdf52d12cd2 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Mon, 6 Jan 2025 12:25:54 +0000 Subject: [PATCH 178/189] chore(Topology/UniformSpace/Completion): make coe' argument implicit (#20514) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Right now coercions to completions of uniform spaces print weirdly: ```lean import Mathlib.Topology.UniformSpace.Completion variable (α : Type) [UniformSpace α] (x : α) open UniformSpace #check (x : Completion α) -- ↑α x : Completion α ``` This is because the coercion from `α` to `Completion α` is `UniformSpace.Completion.coe'` which (accidentally?) has `α` explicit. We make it implicit, so now we get the expected ```lean #check (x : Completion α) -- ↑x : Completion α ``` Zulip discussion: https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/notation.20for.20coercion.20to.20completion/near/439560582 --- Mathlib/Topology/UniformSpace/Completion.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Topology/UniformSpace/Completion.lean b/Mathlib/Topology/UniformSpace/Completion.lean index 7e1f98e7b1de7..d7cc641b4699f 100644 --- a/Mathlib/Topology/UniformSpace/Completion.lean +++ b/Mathlib/Topology/UniformSpace/Completion.lean @@ -310,12 +310,13 @@ instance completeSpace : CompleteSpace (Completion α) := instance t0Space : T0Space (Completion α) := SeparationQuotient.instT0Space +variable {α} in /-- The map from a uniform space to its completion. -/ @[coe] def coe' : α → Completion α := SeparationQuotient.mk ∘ pureCauchy /-- Automatic coercion from `α` to its completion. Not always injective. -/ instance : Coe α (Completion α) := - ⟨coe' α⟩ + ⟨coe'⟩ -- note [use has_coe_t] protected theorem coe_eq : ((↑) : α → Completion α) = SeparationQuotient.mk ∘ pureCauchy := rfl From 88c4a940ec2fc143fcabf33ab1a7f0e56a45d5c6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 6 Jan 2025 12:46:46 +0000 Subject: [PATCH 179/189] chore(Data/Multiset/Basic): move the `sub`, `union`, `inter` sections lower (#19863) * Move the sections about `sub`, `union` and `inter` under the sections about `count` and `filter` * `Multiset.countP_sub` now has a different argument order to match `List.countP_diff` * Golf `List.count_sub` * Name the instances and declare them using `where` notation * Define `Multiset.card` using `Quotient.lift`, not `Quotient.liftOn` --- Mathlib/Data/Multiset/Basic.lean | 629 +++++++++++++++---------------- 1 file changed, 294 insertions(+), 335 deletions(-) diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index f7425d084e924..9fec51e78b073 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -6,6 +6,7 @@ Authors: Mario Carneiro import Mathlib.Algebra.Group.Hom.Defs import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Algebra.Order.Sub.Unbundled.Basic +import Mathlib.Data.List.Count import Mathlib.Data.List.Perm.Basic import Mathlib.Data.List.Perm.Lattice import Mathlib.Data.List.Perm.Subperm @@ -586,6 +587,7 @@ end /-! ### Additive monoid -/ +section add /-- The sum of two multisets is the lift of the list append operation. This adds the multiplicities of each element, @@ -651,6 +653,8 @@ theorem add_cons (a : α) (s t : Multiset α) : s + a ::ₘ t = a ::ₘ (s + t) theorem mem_add {a : α} {s t : Multiset α} : a ∈ s + t ↔ a ∈ s ∨ a ∈ t := Quotient.inductionOn₂ s t fun _l₁ _l₂ => mem_append +end add + theorem mem_of_mem_nsmul {a : α} {s : Multiset α} {n : ℕ} (h : a ∈ n • s) : a ∈ s := by induction' n with n ih · rw [zero_nsmul] at h @@ -679,7 +683,7 @@ theorem nsmul_cons {s : Multiset α} (n : ℕ) (a : α) : /-- The cardinality of a multiset is the sum of the multiplicities of all its elements, or simply the length of the underlying list. -/ -def card (s : Multiset α) : ℕ := Quot.liftOn s length fun _l₁ _l₂ => Perm.length_eq +def card : Multiset α → ℕ := Quot.lift length fun _l₁ _l₂ => Perm.length_eq @[simp] theorem coe_card (l : List α) : card (l : Multiset α) = length l := @@ -1463,260 +1467,6 @@ instance decidableDexistsMultiset {p : ∀ a ∈ m, Prop} [_hp : ∀ (a) (h : a end DecidablePiExists -/-! ### Subtraction -/ - - -section - -variable [DecidableEq α] {s t u : Multiset α} {a : α} - -/-- `s - t` is the multiset such that `count a (s - t) = count a s - count a t` for all `a` - (note that it is truncated subtraction, so it is `0` if `count a t ≥ count a s`). -/ -protected def sub (s t : Multiset α) : Multiset α := - (Quotient.liftOn₂ s t fun l₁ l₂ => (l₁.diff l₂ : Multiset α)) fun _v₁ _v₂ _w₁ _w₂ p₁ p₂ => - Quot.sound <| p₁.diff p₂ - -instance : Sub (Multiset α) := - ⟨Multiset.sub⟩ - -@[simp] -theorem coe_sub (s t : List α) : (s - t : Multiset α) = (s.diff t : List α) := - rfl - -/-- This is a special case of `tsub_zero`, which should be used instead of this. - This is needed to prove `OrderedSub (Multiset α)`. -/ -protected theorem sub_zero (s : Multiset α) : s - 0 = s := - Quot.inductionOn s fun _l => rfl - -@[simp] -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 - revert s - 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] - -theorem sub_eq_fold_erase (s t : Multiset α) : s - t = foldl erase s t := - Quotient.inductionOn₂ s t fun l₁ l₂ => by - show ofList (l₁.diff l₂) = foldl erase l₁ l₂ - rw [diff_eq_foldl l₁ l₂] - symm - exact foldl_hom _ _ _ _ _ fun x y => rfl - -@[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] - -/-! ### Union -/ - - -/-- `s ∪ t` is the lattice join operation with respect to the - multiset `≤`. The multiplicity of `a` in `s ∪ t` is the maximum - of the multiplicities in `s` and `t`. -/ -def union (s t : Multiset α) : Multiset α := - s - t + t - -instance : Union (Multiset α) := - ⟨union⟩ - -theorem union_def (s t : Multiset α) : s ∪ t = s - t + t := - rfl - -theorem le_union_left {s t : Multiset α} : s ≤ s ∪ t := - le_tsub_add - -theorem le_union_right {s t : Multiset α} : t ≤ s ∪ t := - le_add_left _ _ - -theorem eq_union_left : t ≤ s → s ∪ t = s := - tsub_add_cancel_of_le - -@[gcongr] -theorem union_le_union_right (h : s ≤ t) (u) : s ∪ u ≤ t ∪ u := - add_le_add_right (tsub_le_tsub_right h _) 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 <| Multiset.sub_le_self _ _), - (Or.elim · (mem_of_le le_union_left) (mem_of_le le_union_right))⟩ - -@[simp] -theorem map_union [DecidableEq β] {f : α → β} (finj : Function.Injective f) {s t : Multiset α} : - map f (s ∪ t) = map f s ∪ map f t := - Quotient.inductionOn₂ s t fun l₁ l₂ => - congr_arg ofList (by rw [List.map_append f, List.map_diff finj]) - -@[simp] theorem zero_union : 0 ∪ s = s := by - simp [union_def, Multiset.zero_sub] - -@[simp] theorem union_zero : s ∪ 0 = s := by - simp [union_def] - -/-! ### Intersection -/ - -/-- `s ∩ t` is the lattice meet operation with respect to the - multiset `≤`. The multiplicity of `a` in `s ∩ t` is the minimum - of the multiplicities in `s` and `t`. -/ -def inter (s t : Multiset α) : Multiset α := - Quotient.liftOn₂ s t (fun l₁ l₂ => (l₁.bagInter l₂ : Multiset α)) fun _v₁ _v₂ _w₁ _w₂ p₁ p₂ => - Quot.sound <| p₁.bagInter p₂ - -instance : Inter (Multiset α) := - ⟨inter⟩ - -@[simp] -theorem inter_zero (s : Multiset α) : s ∩ 0 = 0 := - Quot.inductionOn s fun l => congr_arg ofList l.bagInter_nil - -@[simp] -theorem zero_inter (s : Multiset α) : 0 ∩ s = 0 := - Quot.inductionOn s fun l => congr_arg ofList l.nil_bagInter - -@[simp] -theorem cons_inter_of_pos {a} (s : Multiset α) {t} : a ∈ t → (a ::ₘ s) ∩ t = a ::ₘ s ∩ t.erase a := - Quotient.inductionOn₂ s t fun _l₁ _l₂ h => congr_arg ofList <| cons_bagInter_of_pos _ h - -@[simp] -theorem cons_inter_of_neg {a} (s : Multiset α) {t} : a ∉ t → (a ::ₘ s) ∩ t = s ∩ t := - Quotient.inductionOn₂ s t fun _l₁ _l₂ h => congr_arg ofList <| cons_bagInter_of_neg _ h - -theorem inter_le_left {s t : Multiset α} : s ∩ t ≤ s := - Quotient.inductionOn₂ s t fun _l₁ _l₂ => (bagInter_sublist_left _ _).subperm - -theorem inter_le_right {s t : Multiset α} : s ∩ t ≤ t := by - induction' s using Multiset.induction_on with a s IH generalizing t - · exact (zero_inter t).symm ▸ zero_le _ - by_cases h : a ∈ t - · simpa [h] using cons_le_cons a (IH (t := t.erase a)) - · simp [h, IH] - -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] 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₂) - · rw [cons_inter_of_neg _ h] - exact IH ((le_cons_of_not_mem <| mt (mem_of_le h₂) h).1 h₁) h₂ - -@[simp] -theorem mem_inter : a ∈ s ∩ t ↔ a ∈ s ∧ a ∈ t := - ⟨fun h => ⟨mem_of_le inter_le_left h, mem_of_le inter_le_right h⟩, fun ⟨h₁, h₂⟩ => by - rw [← cons_erase h₁, cons_inter_of_pos _ h₂]; apply mem_cons_self⟩ - -instance : Lattice (Multiset α) where - sup := (· ∪ ·) - sup_le _ _ _ := union_le - le_sup_left _ _ := le_union_left - le_sup_right _ _ := le_union_right - inf := (· ∩ ·) - le_inf _ _ _ := le_inter - inf_le_left _ _ := inter_le_left - inf_le_right _ _ := inter_le_right - -@[simp] -theorem sup_eq_union (s t : Multiset α) : s ⊔ t = s ∪ t := - rfl - -@[simp] -theorem inf_eq_inter (s t : Multiset α) : s ⊓ t = s ∩ t := - rfl - -@[simp] -theorem le_inter_iff : s ≤ t ∩ u ↔ s ≤ t ∧ s ≤ u := - le_inf_iff - -@[simp] -theorem union_le_iff : s ∪ t ≤ u ↔ s ≤ u ∧ t ≤ u := - sup_le_iff - -theorem union_comm (s t : Multiset α) : s ∪ t = t ∪ s := sup_comm _ _ - -theorem inter_comm (s t : Multiset α) : s ∩ t = t ∩ s := inf_comm _ _ - -theorem eq_union_right (h : s ≤ t) : s ∪ t = t := by rw [union_comm, eq_union_left h] - -@[gcongr] -theorem union_le_union_left (h : s ≤ t) (u) : u ∪ s ≤ u ∪ t := - sup_le_sup_left h _ - -theorem union_le_add (s t : Multiset α) : s ∪ t ≤ s + t := - union_le (le_add_right _ _) (le_add_left _ _) - -theorem union_add_distrib (s t u : Multiset α) : s ∪ t + u = s + u ∪ (t + u) := by - simpa [(· ∪ ·), union, eq_comm, add_assoc] using - show s + u - (t + u) = s - t by rw [add_comm t, tsub_add_eq_tsub_tsub, add_tsub_cancel_right] - -theorem add_union_distrib (s t u : Multiset α) : s + (t ∪ u) = s + t ∪ (s + u) := by - rw [add_comm, union_add_distrib, add_comm s, add_comm s] - -theorem cons_union_distrib (a : α) (s t : Multiset α) : a ::ₘ (s ∪ t) = a ::ₘ s ∪ a ::ₘ t := by - simpa using add_union_distrib (a ::ₘ 0) s t - -theorem inter_add_distrib (s t u : Multiset α) : s ∩ t + u = (s + u) ∩ (t + u) := by - by_contra! h - obtain ⟨a, ha⟩ := lt_iff_cons_le.1 <| h.lt_of_le <| le_inter - (add_le_add_right inter_le_left _) (add_le_add_right inter_le_right _) - rw [← cons_add] at ha - exact (lt_cons_self (s ∩ t) a).not_le <| le_inter - (le_of_add_le_add_right (ha.trans inter_le_left)) - (le_of_add_le_add_right (ha.trans inter_le_right)) - -theorem add_inter_distrib (s t u : Multiset α) : s + t ∩ u = (s + t) ∩ (s + u) := by - rw [add_comm, inter_add_distrib, add_comm s, add_comm s] - -theorem cons_inter_distrib (a : α) (s t : Multiset α) : a ::ₘ s ∩ t = (a ::ₘ s) ∩ (a ::ₘ t) := by - simp - -lemma union_add_inter (s t : Multiset α) : s ∪ t + s ∩ t = s + t := by - apply _root_.le_antisymm - · rw [union_add_distrib] - refine union_le (add_le_add_left inter_le_right _) ?_ - rw [add_comm] - exact add_le_add_right inter_le_left _ - · rw [add_comm, add_inter_distrib] - refine le_inter (add_le_add_right le_union_right _) ?_ - rw [add_comm] - exact add_le_add_right le_union_left _ - -theorem sub_add_inter (s t : Multiset α) : s - t + s ∩ t = s := by - rw [inter_comm] - revert s; refine Multiset.induction_on t (by simp) fun a t IH s => ?_ - by_cases h : a ∈ s - · rw [cons_inter_of_pos _ h, sub_cons, add_cons, IH, cons_erase h] - · rw [cons_inter_of_neg _ h, sub_cons, erase_of_not_mem h, IH] - -theorem sub_inter (s t : Multiset α) : s - s ∩ t = s - t := - add_right_cancel (b := s ∩ t) <| by - rw [sub_add_inter s t, tsub_add_cancel_of_le inter_le_left] - -end - /-! ### `Multiset.filter` -/ @@ -1822,35 +1572,6 @@ theorem filter_nsmul (s : Multiset α) (n : ℕ) : filter p (n • s) = n • fi variable (p) -@[simp] -theorem filter_sub [DecidableEq α] (s t : Multiset α) : - filter p (s - t) = filter p s - filter p t := by - revert s; refine Multiset.induction_on t (by simp) fun a t IH s => ?_ - rw [sub_cons, IH] - by_cases h : p a - · rw [filter_cons_of_pos _ h, sub_cons] - congr - by_cases m : a ∈ s - · rw [← cons_inj_right a, ← filter_cons_of_pos _ h, cons_erase (mem_filter_of_mem m h), - cons_erase m] - · rw [erase_of_not_mem m, erase_of_not_mem (mt mem_of_mem_filter m)] - · rw [filter_cons_of_neg _ h] - by_cases m : a ∈ s - · rw [(by rw [filter_cons_of_neg _ h] : filter p (erase s a) = filter p (a ::ₘ erase s a)), - cons_erase m] - · rw [erase_of_not_mem m] - -@[simp] -theorem filter_union [DecidableEq α] (s t : Multiset α) : - filter p (s ∪ t) = filter p s ∪ filter p t := by simp [(· ∪ ·), union] - -@[simp] -theorem filter_inter [DecidableEq α] (s t : Multiset α) : - filter p (s ∩ t) = filter p s ∩ filter p t := - le_antisymm (le_inter (filter_le_filter _ inter_le_left) (filter_le_filter _ inter_le_right)) <| - le_filter.2 ⟨inf_le_inf (filter_le _ _) (filter_le _ _), fun _a h => - of_mem_filter (mem_of_le inter_le_left h)⟩ - @[simp] theorem filter_filter (q) [DecidablePred q] (s : Multiset α) : filter p (filter q s) = filter (fun a => p a ∧ q a) s := @@ -2026,11 +1747,6 @@ def countPAddMonoidHom : Multiset α →+ ℕ where theorem coe_countPAddMonoidHom : (countPAddMonoidHom p : Multiset α → ℕ) = countP p := rfl -@[simp] -theorem countP_sub [DecidableEq α] {s t : Multiset α} (h : t ≤ s) : - countP p (s - t) = countP p s - countP p t := by - simp [countP_eq_card_filter, h, filter_le_filter] - @[gcongr] theorem countP_le_of_le {s t} (h : s ≤ t) : countP p s ≤ countP p t := by simpa [countP_eq_card_filter] using card_le_card (filter_le_filter p h) @@ -2209,23 +1925,6 @@ theorem count_erase_of_ne {a b : α} (ab : a ≠ b) (s : Multiset α) : Quotient.inductionOn s fun l => by convert List.count_erase_of_ne ab l <;> rw [← coe_count] <;> simp -@[simp] -theorem count_sub (a : α) (s t : Multiset α) : count a (s - t) = count a s - count a t := by - revert s; refine Multiset.induction_on t (by simp) fun b t IH s => ?_ - rw [sub_cons, IH] - rcases Decidable.eq_or_ne a b with rfl | ab - · rw [count_erase_self, count_cons_self, Nat.sub_sub, add_comm] - · rw [count_erase_of_ne ab, count_cons_of_ne ab] - -@[simp] -theorem count_union (a : α) (s t : Multiset α) : count a (s ∪ t) = max (count a s) (count a t) := by - simp [(· ∪ ·), union, Nat.sub_add_eq_max] - -@[simp] -theorem count_inter (a : α) (s t : Multiset α) : count a (s ∩ t) = min (count a s) (count a t) := by - apply @Nat.add_left_cancel (count a (s - t)) - rw [← count_add, sub_add_inter, count_sub, Nat.sub_add_min_cancel] - theorem le_count_iff_replicate_le {a : α} {s : Multiset α} {n : ℕ} : n ≤ count a s ↔ replicate n a ≤ s := Quot.inductionOn s fun _l => by @@ -2263,20 +1962,8 @@ theorem ext' {s t : Multiset α} : (∀ a, count a s = count a t) → s = t := lemma count_injective : Injective fun (s : Multiset α) a ↦ s.count a := 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 - theorem le_iff_count {s t : Multiset α} : s ≤ t ↔ ∀ a, count a s ≤ count a t := - ⟨fun h a => count_le_of_le a h, fun al => by - rw [← (ext.2 fun a => by simp [max_eq_right (al a)] : s ∪ t = t)]; apply le_union_left⟩ - -instance : DistribLattice (Multiset α) := - { le_sup_inf := fun s t u => - le_of_eq <| - Eq.symm <| - ext.2 fun a => by - simp only [max_min_distrib_left, Multiset.count_inter, Multiset.sup_eq_union, - Multiset.count_union, Multiset.inf_eq_inter] } + Quotient.inductionOn₂ s t fun _ _ ↦ by simp [subperm_iff_count] theorem count_map {α β : Type*} (f : α → β) (s : Multiset α) [DecidableEq β] (b : β) : count b (map f s) = card (s.filter fun a => b = f a) := by @@ -2303,11 +1990,6 @@ theorem count_map_eq_count' [DecidableEq β] (f : α → β) (s : Multiset α) ( rw [hf hkx] at hks contradiction -@[simp] -theorem sub_filter_eq_filter_not (p) [DecidablePred p] (s : Multiset α) : - s - s.filter p = s.filter (fun a ↦ ¬ p a) := by - ext a; by_cases h : p a <;> simp [h] - theorem filter_eq' (s : Multiset α) (b : α) : s.filter (· = b) = replicate (count b s) b := Quotient.inductionOn s fun l => by simp only [quot_mk_to_coe, filter_coe, mem_coe, coe_count] @@ -2316,6 +1998,292 @@ theorem filter_eq' (s : Multiset α) (b : α) : s.filter (· = b) = replicate (c theorem filter_eq (s : Multiset α) (b : α) : s.filter (Eq b) = replicate (count b s) b := by simp_rw [← filter_eq', eq_comm] +lemma erase_attach_map_val (s : Multiset α) (x : {x // x ∈ s}) : + (s.attach.erase x).map (↑) = s.erase x := by + rw [Multiset.map_erase _ val_injective, attach_map_val] + +lemma erase_attach_map (s : Multiset α) (f : α → β) (x : {x // x ∈ s}) : + (s.attach.erase x).map (fun j : {x // x ∈ s} ↦ f j) = (s.erase x).map f := by + simp only [← Function.comp_apply (f := f)] + rw [← map_map, erase_attach_map_val] + +end + +/-! ### Subtraction -/ + +section sub +variable [DecidableEq α] {s t u : Multiset α} {a : α} + +/-- `s - t` is the multiset such that `count a (s - t) = count a s - count a t` for all `a`. +(note that it is truncated subtraction, so `count a (s - t) = 0` if `count a s ≤ count a t`). -/ +protected def sub (s t : Multiset α) : Multiset α := + (Quotient.liftOn₂ s t fun l₁ l₂ => (l₁.diff l₂ : Multiset α)) fun _v₁ _v₂ _w₁ _w₂ p₁ p₂ => + Quot.sound <| p₁.diff p₂ + +instance : Sub (Multiset α) := ⟨.sub⟩ + +@[simp] +lemma coe_sub (s t : List α) : (s - t : Multiset α) = s.diff t := + rfl + +/-- This is a special case of `tsub_zero`, which should be used instead of this. +This is needed to prove `OrderedSub (Multiset α)`. -/ +@[simp, nolint simpNF] -- We want to use this lemma earlier than the lemma simp can prove it with +protected lemma sub_zero (s : Multiset α) : s - 0 = s := + Quot.inductionOn s fun _l => rfl + +@[simp] +lemma 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 lemma zero_sub (t : Multiset α) : 0 - t = 0 := + Multiset.induction_on t rfl fun a s ih => by simp [ih] + +@[simp] +lemma countP_sub {s t : Multiset α} : + t ≤ s → ∀ (p : α → Prop) [DecidablePred p], countP p (s - t) = countP p s - countP p t := + Quotient.inductionOn₂ s t fun _l₁ _l₂ hl _ _ ↦ List.countP_diff hl _ + +@[simp] +lemma count_sub (a : α) (s t : Multiset α) : count a (s - t) = count a s - count a t := + Quotient.inductionOn₂ s t <| by simp [List.count_diff] + +/-- 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 lemma sub_le_iff_le_add : s - t ≤ u ↔ s ≤ u + t := by + revert s + exact @(Multiset.induction_on t (by simp [Multiset.sub_zero]) fun a t IH s => by + simp [IH, erase_le_iff_le_cons]) + +protected lemma 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⟩ + +@[simp] +lemma filter_sub (p : α → Prop) [DecidablePred p] (s t : Multiset α) : + filter p (s - t) = filter p s - filter p t := by + revert s; refine Multiset.induction_on t (by simp) fun a t IH s => ?_ + rw [sub_cons, IH] + by_cases h : p a + · rw [filter_cons_of_pos _ h, sub_cons] + congr + by_cases m : a ∈ s + · rw [← cons_inj_right a, ← filter_cons_of_pos _ h, cons_erase (mem_filter_of_mem m h), + cons_erase m] + · rw [erase_of_not_mem m, erase_of_not_mem (mt mem_of_mem_filter m)] + · rw [filter_cons_of_neg _ h] + by_cases m : a ∈ s + · rw [(by rw [filter_cons_of_neg _ h] : filter p (erase s a) = filter p (a ::ₘ erase s a)), + cons_erase m] + · rw [erase_of_not_mem m] + +@[simp] +lemma sub_filter_eq_filter_not (p : α → Prop) [DecidablePred p] (s : Multiset α) : + s - s.filter p = s.filter fun a ↦ ¬ p a := by ext a; by_cases h : p a <;> simp [h] + +lemma 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] + +lemma sub_eq_fold_erase (s t : Multiset α) : s - t = foldl erase s t := + Quotient.inductionOn₂ s t fun l₁ l₂ => by + show ofList (l₁.diff l₂) = foldl erase l₁ l₂ + rw [diff_eq_foldl l₁ l₂] + symm + exact foldl_hom _ _ _ _ _ fun x y => rfl + +@[simp] +lemma 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] + +/-! ### Union -/ + +/-- `s ∪ t` is the multiset such that the multiplicity of each `a` in it is the maximum of the +multiplicity of `a` in `s` and `t`. This is the supremum of multisets. -/ +def union (s t : Multiset α) : Multiset α := s - t + t + +instance : Union (Multiset α) := ⟨union⟩ + +lemma union_def (s t : Multiset α) : s ∪ t = s - t + t := rfl + +lemma le_union_left : s ≤ s ∪ t := le_tsub_add +lemma le_union_right : t ≤ s ∪ t := le_add_left _ _ +lemma eq_union_left : t ≤ s → s ∪ t = s := tsub_add_cancel_of_le + +@[gcongr] +lemma union_le_union_right (h : s ≤ t) (u) : s ∪ u ≤ t ∪ u := + add_le_add_right (tsub_le_tsub_right h _) _ + +lemma 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] +lemma mem_union : a ∈ s ∪ t ↔ a ∈ s ∨ a ∈ t := + ⟨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] +lemma map_union [DecidableEq β] {f : α → β} (finj : Function.Injective f) {s t : Multiset α} : + map f (s ∪ t) = map f s ∪ map f t := + Quotient.inductionOn₂ s t fun l₁ l₂ => + congr_arg ofList (by rw [List.map_append f, List.map_diff finj]) + +@[simp] lemma zero_union : 0 ∪ s = s := by simp [union_def, Multiset.zero_sub] +@[simp] lemma union_zero : s ∪ 0 = s := by simp [union_def] + +@[simp] +lemma count_union (a : α) (s t : Multiset α) : count a (s ∪ t) = max (count a s) (count a t) := by + simp [(· ∪ ·), union, Nat.sub_add_eq_max] + +@[simp] lemma filter_union (p : α → Prop) [DecidablePred p] (s t : Multiset α) : + filter p (s ∪ t) = filter p s ∪ filter p t := by simp [(· ∪ ·), union] + +/-! ### Intersection -/ + +/-- `s ∩ t` is the multiset such that the multiplicity of each `a` in it is the minimum of the +multiplicity of `a` in `s` and `t`. This is the infimum of multisets. -/ +def inter (s t : Multiset α) : Multiset α := + Quotient.liftOn₂ s t (fun l₁ l₂ => (l₁.bagInter l₂ : Multiset α)) fun _v₁ _v₂ _w₁ _w₂ p₁ p₂ => + Quot.sound <| p₁.bagInter p₂ + +instance : Inter (Multiset α) := ⟨inter⟩ + +@[simp] lemma inter_zero (s : Multiset α) : s ∩ 0 = 0 := + Quot.inductionOn s fun l => congr_arg ofList l.bagInter_nil + +@[simp] lemma zero_inter (s : Multiset α) : 0 ∩ s = 0 := + Quot.inductionOn s fun l => congr_arg ofList l.nil_bagInter + +@[simp] +lemma cons_inter_of_pos (s : Multiset α) : a ∈ t → (a ::ₘ s) ∩ t = a ::ₘ s ∩ t.erase a := + Quotient.inductionOn₂ s t fun _l₁ _l₂ h => congr_arg ofList <| cons_bagInter_of_pos _ h + +@[simp] +lemma cons_inter_of_neg (s : Multiset α) : a ∉ t → (a ::ₘ s) ∩ t = s ∩ t := + Quotient.inductionOn₂ s t fun _l₁ _l₂ h => congr_arg ofList <| cons_bagInter_of_neg _ h + +lemma inter_le_left : s ∩ t ≤ s := + Quotient.inductionOn₂ s t fun _l₁ _l₂ => (bagInter_sublist_left _ _).subperm + +lemma inter_le_right : s ∩ t ≤ t := by + induction' s using Multiset.induction_on with a s IH generalizing t + · exact (zero_inter t).symm ▸ zero_le _ + by_cases h : a ∈ t + · simpa [h] using cons_le_cons a (IH (t := t.erase a)) + · simp [h, IH] + +lemma 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] 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₂) + · rw [cons_inter_of_neg _ h] + exact IH ((le_cons_of_not_mem <| mt (mem_of_le h₂) h).1 h₁) h₂ + +@[simp] +lemma mem_inter : a ∈ s ∩ t ↔ a ∈ s ∧ a ∈ t := + ⟨fun h => ⟨mem_of_le inter_le_left h, mem_of_le inter_le_right h⟩, fun ⟨h₁, h₂⟩ => by + rw [← cons_erase h₁, cons_inter_of_pos _ h₂]; apply mem_cons_self⟩ + +instance instLattice : Lattice (Multiset α) where + sup := (· ∪ ·) + sup_le _ _ _ := union_le + le_sup_left _ _ := le_union_left + le_sup_right _ _ := le_union_right + inf := (· ∩ ·) + le_inf _ _ _ := le_inter + inf_le_left _ _ := inter_le_left + inf_le_right _ _ := inter_le_right + +@[simp] lemma sup_eq_union (s t : Multiset α) : s ⊔ t = s ∪ t := rfl +@[simp] lemma inf_eq_inter (s t : Multiset α) : s ⊓ t = s ∩ t := rfl + +@[simp] lemma le_inter_iff : s ≤ t ∩ u ↔ s ≤ t ∧ s ≤ u := le_inf_iff +@[simp] lemma union_le_iff : s ∪ t ≤ u ↔ s ≤ u ∧ t ≤ u := sup_le_iff + +lemma union_comm (s t : Multiset α) : s ∪ t = t ∪ s := sup_comm .. +lemma inter_comm (s t : Multiset α) : s ∩ t = t ∩ s := inf_comm .. + +lemma eq_union_right (h : s ≤ t) : s ∪ t = t := by rw [union_comm, eq_union_left h] + +@[gcongr] lemma union_le_union_left (h : s ≤ t) (u) : u ∪ s ≤ u ∪ t := sup_le_sup_left h _ + +lemma union_le_add (s t : Multiset α) : s ∪ t ≤ s + t := union_le (le_add_right ..) (le_add_left ..) + +lemma union_add_distrib (s t u : Multiset α) : s ∪ t + u = s + u ∪ (t + u) := by + simpa [(· ∪ ·), union, eq_comm, add_assoc, add_left_inj] using + show s + u - (t + u) = s - t by + rw [add_comm t, tsub_add_eq_tsub_tsub, add_tsub_cancel_right] + +lemma add_union_distrib (s t u : Multiset α) : s + (t ∪ u) = s + t ∪ (s + u) := by + rw [add_comm, union_add_distrib, add_comm s, add_comm s] + +lemma cons_union_distrib (a : α) (s t : Multiset α) : a ::ₘ (s ∪ t) = a ::ₘ s ∪ a ::ₘ t := by + simpa using add_union_distrib (a ::ₘ 0) s t + +lemma inter_add_distrib (s t u : Multiset α) : s ∩ t + u = (s + u) ∩ (t + u) := by + by_contra! h + obtain ⟨a, ha⟩ := lt_iff_cons_le.1 <| h.lt_of_le <| le_inter + (add_le_add_right inter_le_left _) (add_le_add_right inter_le_right _) + rw [← cons_add] at ha + exact (lt_cons_self (s ∩ t) a).not_le <| le_inter + (le_of_add_le_add_right (ha.trans inter_le_left)) + (le_of_add_le_add_right (ha.trans inter_le_right)) + +lemma add_inter_distrib (s t u : Multiset α) : s + t ∩ u = (s + t) ∩ (s + u) := by + rw [add_comm, inter_add_distrib, add_comm s, add_comm s] + +lemma cons_inter_distrib (a : α) (s t : Multiset α) : a ::ₘ s ∩ t = (a ::ₘ s) ∩ (a ::ₘ t) := by + simp + +lemma union_add_inter (s t : Multiset α) : s ∪ t + s ∩ t = s + t := by + apply _root_.le_antisymm + · rw [union_add_distrib] + refine union_le (add_le_add_left inter_le_right _) ?_ + rw [add_comm] + exact add_le_add_right inter_le_left _ + · rw [add_comm, add_inter_distrib] + refine le_inter (add_le_add_right le_union_right _) ?_ + rw [add_comm] + exact add_le_add_right le_union_left _ + +lemma sub_add_inter (s t : Multiset α) : s - t + s ∩ t = s := by + rw [inter_comm] + revert s; refine Multiset.induction_on t (by simp) fun a t IH s => ?_ + by_cases h : a ∈ s + · rw [cons_inter_of_pos _ h, sub_cons, add_cons, IH, cons_erase h] + · rw [cons_inter_of_neg _ h, sub_cons, erase_of_not_mem h, IH] + +lemma sub_inter (s t : Multiset α) : s - s ∩ t = s - t := + (eq_tsub_of_add_eq <| sub_add_inter ..).symm + +@[simp] +lemma count_inter (a : α) (s t : Multiset α) : count a (s ∩ t) = min (count a s) (count a t) := by + apply @Nat.add_left_cancel (count a (s - t)) + rw [← count_add, sub_add_inter, count_sub, Nat.sub_add_min_cancel] + +@[simp] +lemma coe_inter (s t : List α) : (s ∩ t : Multiset α) = (s.bagInter t : List α) := by ext; simp + +instance instDistribLattice : DistribLattice (Multiset α) where + le_sup_inf s t u := ge_of_eq <| ext.2 fun a ↦ by + simp only [max_min_distrib_left, Multiset.count_inter, Multiset.sup_eq_union, + Multiset.count_union, Multiset.inf_eq_inter] + +@[simp] lemma filter_inter (p : α → Prop) [DecidablePred p] (s t : Multiset α) : + filter p (s ∩ t) = filter p s ∩ filter p t := + le_antisymm (le_inter (filter_le_filter _ inter_le_left) (filter_le_filter _ inter_le_right)) <| + le_filter.2 ⟨inf_le_inf (filter_le _ _) (filter_le _ _), fun _a h => + of_mem_filter (mem_of_le inter_le_left h)⟩ + @[simp] theorem replicate_inter (n : ℕ) (x : α) (s : Multiset α) : replicate n x ∩ s = replicate (min n (s.count x)) x := by @@ -2330,16 +2298,7 @@ theorem inter_replicate (s : Multiset α) (n : ℕ) (x : α) : s ∩ replicate n x = replicate (min (s.count x) n) x := by rw [inter_comm, replicate_inter, min_comm] -theorem erase_attach_map_val (s : Multiset α) (x : {x // x ∈ s}) : - (s.attach.erase x).map (↑) = s.erase x := by - rw [Multiset.map_erase _ val_injective, attach_map_val] - -theorem erase_attach_map (s : Multiset α) (f : α → β) (x : {x // x ∈ s}) : - (s.attach.erase x).map (fun j : {x // x ∈ s} ↦ f j) = (s.erase x).map f := by - simp only [← Function.comp_apply (f := f)] - rw [← map_map, erase_attach_map_val] - -end +end sub @[ext] theorem addHom_ext [AddZeroClass β] ⦃f g : Multiset α →+ β⦄ (h : ∀ x, f {x} = g {x}) : f = g := by @@ -2800,4 +2759,4 @@ theorem coe_subsingletonEquiv [Subsingleton α] : end Multiset -set_option linter.style.longFile 2900 +set_option linter.style.longFile 2800 From 5e298f933b82bbc911d081bf7b21b1e5d7a24ae9 Mon Sep 17 00:00:00 2001 From: smorel394 Date: Mon, 6 Jan 2025 13:13:30 +0000 Subject: [PATCH 180/189] feat(CategoryTheory/Shift/Adjunction): properties of `Adjunction.CommShift` (#20337) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Provide instances that say that `Adjunction.CommShift` holds for the identity adjunction and is stable by composition, as well as similar instances for `Equivalence.CommShift`. - [ ] depends on: #20364 - [ ] depends on: #20490 Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> Co-authored-by: smorel394 <67864981+smorel394@users.noreply.github.com> --- Mathlib/CategoryTheory/Shift/Adjunction.lean | 102 +++++++++++++++++-- 1 file changed, 96 insertions(+), 6 deletions(-) diff --git a/Mathlib/CategoryTheory/Shift/Adjunction.lean b/Mathlib/CategoryTheory/Shift/Adjunction.lean index f893e59286f66..8f6a5a6ddc1f9 100644 --- a/Mathlib/CategoryTheory/Shift/Adjunction.lean +++ b/Mathlib/CategoryTheory/Shift/Adjunction.lean @@ -200,9 +200,31 @@ class CommShift : Prop where commShift_unit : NatTrans.CommShift adj.unit A := by infer_instance commShift_counit : NatTrans.CommShift adj.counit A := by infer_instance +open CommShift in +attribute [instance] commShift_unit commShift_counit + +@[reassoc (attr := simp)] +lemma unit_app_commShiftIso_hom_app [adj.CommShift A] (a : A) (X : C) : + adj.unit.app (X⟦a⟧) ≫ ((F ⋙ G).commShiftIso a).hom.app X = (adj.unit.app X)⟦a⟧' := by + simpa using (NatTrans.shift_app_comm adj.unit a X).symm + +@[reassoc (attr := simp)] +lemma unit_app_shift_commShiftIso_inv_app [adj.CommShift A] (a : A) (X : C) : + (adj.unit.app X)⟦a⟧' ≫ ((F ⋙ G).commShiftIso a).inv.app X = adj.unit.app (X⟦a⟧) := by + simp [← cancel_mono (((F ⋙ G).commShiftIso _).hom.app _)] + +@[reassoc (attr := simp)] +lemma commShiftIso_hom_app_counit_app_shift [adj.CommShift A] (a : A) (Y : D) : + ((G ⋙ F).commShiftIso a).hom.app Y ≫ (adj.counit.app Y)⟦a⟧' = adj.counit.app (Y⟦a⟧) := by + simpa using (NatTrans.shift_app_comm adj.counit a Y) + +@[reassoc (attr := simp)] +lemma commShiftIso_inv_app_counit_app [adj.CommShift A] (a : A) (Y : D) : + ((G ⋙ F).commShiftIso a).inv.app Y ≫ adj.counit.app (Y⟦a⟧) = (adj.counit.app Y)⟦a⟧' := by + simp [← cancel_epi (((G ⋙ F).commShiftIso _).hom.app _)] + namespace CommShift -attribute [instance] commShift_unit commShift_counit /-- Constructor for `Adjunction.CommShift`. -/ lemma mk' (h : NatTrans.CommShift adj.unit A) : @@ -217,6 +239,30 @@ lemma mk' (h : NatTrans.CommShift adj.unit A) : Functor.commShiftIso_id_hom_app, whiskerRight_app, id_comp, Functor.commShiftIso_comp_hom_app] using congr_app (h.shift_comm a) X⟩ +variable [adj.CommShift A] + +/-- The identity adjunction is compatible with the trivial `CommShift` structure on the +identity functor. +-/ +instance instId : (Adjunction.id (C := C)).CommShift A where + commShift_counit := + inferInstanceAs (NatTrans.CommShift (𝟭 C).leftUnitor.hom A) + commShift_unit := + inferInstanceAs (NatTrans.CommShift (𝟭 C).leftUnitor.inv A) + +variable {E : Type*} [Category E] {F' : D ⥤ E} {G' : E ⥤ D} (adj' : F' ⊣ G') + [HasShift E A] [F'.CommShift A] [G'.CommShift A] [adj.CommShift A] [adj'.CommShift A] + +/-- Compatibility of `Adjunction.Commshift` with the composition of adjunctions. +-/ +instance instComp : (adj.comp adj').CommShift A where + commShift_counit := by + rw [comp_counit] + infer_instance + commShift_unit := by + rw [comp_unit] + infer_instance + end CommShift variable {A} @@ -288,8 +334,8 @@ lemma iso_inv_app (Y : D) : (F.obj (G.obj Y)))) ≫ G.map ((shiftFunctor D a).map (adj.counit.app Y)) := by obtain rfl : b = -a := by rw [← add_left_inj a, h, neg_add_cancel] - simp only [Functor.comp_obj, iso, iso', shiftEquiv', Equiv.toFun_as_coe, - conjugateIsoEquiv_apply_inv, conjugateEquiv_apply_app, comp_unit_app, Functor.id_obj, + simp only [iso, iso', shiftEquiv', Equiv.toFun_as_coe, conjugateIsoEquiv_apply_inv, + conjugateEquiv_apply_app, Functor.comp_obj, comp_unit_app, Functor.id_obj, Equivalence.toAdjunction_unit, Equivalence.Equivalence_mk'_unit, Iso.symm_hom, Functor.comp_map, comp_counit_app, Equivalence.toAdjunction_counit, Equivalence.Equivalence_mk'_counit, Functor.map_shiftFunctorCompIsoId_hom_app, assoc, Functor.map_comp] @@ -470,9 +516,28 @@ lemma mk' (h : NatTrans.CommShift E.unitIso.hom A) : commShift_counit := (Adjunction.CommShift.mk' E.toAdjunction A h).commShift_counit /-- -If `E : C ≌ D` is an equivalence and we have compatible `CommShift` structures on `E.functor` -and `E.inverse`, then we also have compatible `CommShift` structures on `E.symm.functor` -and `E.symm.inverse`. +The forward functor of the identity equivalence is compatible with shifts. +-/ +instance : (Equivalence.refl (C := C)).functor.CommShift A := by + dsimp + infer_instance + +/-- +The inverse functor of the identity equivalence is compatible with shifts. +-/ +instance : (Equivalence.refl (C := C)).inverse.CommShift A := by + dsimp + infer_instance + +/-- +The identity equivalence is compatible with shifts. +-/ +instance : (Equivalence.refl (C := C)).CommShift A := by + dsimp [Equivalence.CommShift, refl_toAdjunction] + infer_instance + +/-- +If an equivalence `E : C ≌ D` is compatible with shifts, so is `E.symm`. -/ instance [E.CommShift A] : E.symm.CommShift A := mk' E.symm A (inferInstanceAs (NatTrans.CommShift E.counitIso.inv A)) @@ -483,6 +548,31 @@ lemma mk'' (h : NatTrans.CommShift E.counitIso.hom A) : have := mk' E.symm A (inferInstanceAs (NatTrans.CommShift E.counitIso.inv A)) inferInstanceAs (E.symm.symm.CommShift A) +variable {F : Type*} [Category F] [HasShift F A] {E' : D ≌ F} [E.CommShift A] + [E'.functor.CommShift A] [E'.inverse.CommShift A] [E'.CommShift A] + +/-- +If `E : C ≌ D` and `E' : D ≌ F` are equivalence whose forward functors are compatible with shifts, +so is `(E.trans E').functor`. +-/ +instance : (E.trans E').functor.CommShift A := by + dsimp + infer_instance + +/-- +If `E : C ≌ D` and `E' : D ≌ F` are equivalence whose inverse functors are compatible with shifts, +so is `(E.trans E').inverse`. +-/ +instance : (E.trans E').inverse.CommShift A := by + dsimp + infer_instance + +/-- +If equivalences `E : C ≌ D` and `E' : D ≌ F` are compatible with shifts, so is `E.trans E'`. +-/ +instance : (E.trans E').CommShift A := + inferInstanceAs ((E.toAdjunction.comp E'.toAdjunction).CommShift A) + end CommShift end From fceef13c1e9adc7b86abf7215e508050a45bc2ad Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Mon, 6 Jan 2025 13:23:12 +0000 Subject: [PATCH 181/189] fix: Allow multiple builds on staging branch (#20515) as mentioned on [zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/ci.20on.20master.20is.20failing/near/492095011), staging CI fails, which we don't want. This pr allows multiple build runs on the `staging` branch at a time, which should hopefully fix the issue --- .github/build.in.yml | 3 ++- .github/workflows/bors.yml | 3 ++- .github/workflows/build.yml | 3 ++- .github/workflows/build_fork.yml | 3 ++- 4 files changed, 8 insertions(+), 4 deletions(-) diff --git a/.github/build.in.yml b/.github/build.in.yml index b7acc84572050..69c40f7765dae 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -15,7 +15,8 @@ env: concurrency: # label each workflow run; only the latest with each label will run # workflows on master get more expressive labels - group: ${{ github.workflow }}-${{ github.ref }}.${{(github.ref == 'refs/heads/master' && github.run_id) || ''}} + group: ${{ github.workflow }}-${{ github.ref }}. + ${{ ( contains(fromJSON( '["refs/heads/master", "refs/heads/staging"]'), github.ref ) && github.run_id) || ''}} # cancel any running workflow with the same label cancel-in-progress: true diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 37c0ac50d19a0..3149eccc52722 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -25,7 +25,8 @@ env: concurrency: # label each workflow run; only the latest with each label will run # workflows on master get more expressive labels - group: ${{ github.workflow }}-${{ github.ref }}.${{(github.ref == 'refs/heads/master' && github.run_id) || ''}} + group: ${{ github.workflow }}-${{ github.ref }}. + ${{ ( contains(fromJSON( '["refs/heads/master", "refs/heads/staging"]'), github.ref ) && github.run_id) || ''}} # cancel any running workflow with the same label cancel-in-progress: true diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 62525af3c3089..5a3bc5d84c1b0 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -32,7 +32,8 @@ env: concurrency: # label each workflow run; only the latest with each label will run # workflows on master get more expressive labels - group: ${{ github.workflow }}-${{ github.ref }}.${{(github.ref == 'refs/heads/master' && github.run_id) || ''}} + group: ${{ github.workflow }}-${{ github.ref }}. + ${{ ( contains(fromJSON( '["refs/heads/master", "refs/heads/staging"]'), github.ref ) && github.run_id) || ''}} # cancel any running workflow with the same label cancel-in-progress: true diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 14ae2662c9d09..d7ceebe6b73e3 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -29,7 +29,8 @@ env: concurrency: # label each workflow run; only the latest with each label will run # workflows on master get more expressive labels - group: ${{ github.workflow }}-${{ github.ref }}.${{(github.ref == 'refs/heads/master' && github.run_id) || ''}} + group: ${{ github.workflow }}-${{ github.ref }}. + ${{ ( contains(fromJSON( '["refs/heads/master", "refs/heads/staging"]'), github.ref ) && github.run_id) || ''}} # cancel any running workflow with the same label cancel-in-progress: true From d5cc7c85a7795dcbd6fca653498bd3e643edba0a Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Mon, 6 Jan 2025 14:27:26 +0000 Subject: [PATCH 182/189] feat(RingTheory): classification of etale algebras over fields (#20324) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/RingTheory/Etale/Field.lean | 70 +++++++++++++++++++++++++-- Mathlib/RingTheory/Etale/Pi.lean | 36 ++++++++++++++ Mathlib/RingTheory/Smooth/Pi.lean | 3 ++ Mathlib/RingTheory/Unramified/Pi.lean | 3 ++ 5 files changed, 109 insertions(+), 4 deletions(-) create mode 100644 Mathlib/RingTheory/Etale/Pi.lean diff --git a/Mathlib.lean b/Mathlib.lean index 99606d29a64b3..2a2fd484ec0df 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4388,6 +4388,7 @@ import Mathlib.RingTheory.EisensteinCriterion import Mathlib.RingTheory.EssentialFiniteness import Mathlib.RingTheory.Etale.Basic import Mathlib.RingTheory.Etale.Field +import Mathlib.RingTheory.Etale.Pi import Mathlib.RingTheory.EuclideanDomain import Mathlib.RingTheory.Extension import Mathlib.RingTheory.Filtration diff --git a/Mathlib/RingTheory/Etale/Field.lean b/Mathlib/RingTheory/Etale/Field.lean index 3e6b3bd1c7814..be84c58cf9418 100644 --- a/Mathlib/RingTheory/Etale/Field.lean +++ b/Mathlib/RingTheory/Etale/Field.lean @@ -3,7 +3,7 @@ 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.Etale.Basic +import Mathlib.RingTheory.Etale.Pi import Mathlib.RingTheory.Unramified.Field /-! @@ -14,9 +14,14 @@ import Mathlib.RingTheory.Unramified.Field Let `K` be a field, `A` be a `K`-algebra and `L` be a field extension of `K`. - `Algebra.FormallyEtale.of_isSeparable`: - If `L` is separable over `K`, then `L` is formally etale over `K`. + If `L` is separable over `K`, then `L` is formally étale over `K`. - `Algebra.FormallyEtale.iff_isSeparable`: - If `L` is (essentially) of finite type over `K`, then `L/K` is etale iff `L/K` is separable. + If `L` is (essentially) of finite type over `K`, then `L/K` is étale iff `L/K` is separable. +- `Algebra.FormallyEtale.iff_exists_algEquiv_prod`: + If `A` is (essentially) of finite type over `K`, + then `A/K` is étale iff `A` is a finite product of separable field extensions. +- `Algebra.Etale.iff_exists_algEquiv_prod`: + `A/K` is étale iff `A` is a finite product of finite separable field extensions. ## References @@ -27,7 +32,7 @@ Let `K` be a field, `A` be a `K`-algebra and `L` be a field extension of `K`. universe u -variable (K L : Type u) [Field K] [Field L] [Algebra K L] +variable (K L A : Type u) [Field K] [Field L] [CommRing A] [Algebra K L] [Algebra K A] open Algebra Polynomial @@ -148,4 +153,61 @@ theorem iff_isSeparable [EssFiniteType K L] : FormallyEtale K L ↔ Algebra.IsSeparable K L := ⟨fun _ ↦ FormallyUnramified.isSeparable K L, fun _ ↦ of_isSeparable K L⟩ +attribute [local instance] + IsArtinianRing.subtype_isMaximal_finite IsArtinianRing.fieldOfSubtypeIsMaximal in +/-- +If `A` is an essentially of finite type algebra over a field `K`, then `A` is formally étale +over `K` if and only if `A` is a finite product of separable field extensions. +-/ +theorem iff_exists_algEquiv_prod [EssFiniteType K A] : + FormallyEtale K A ↔ + ∃ (I : Type u) (_ : Finite I) (Ai : I → Type u) (_ : ∀ i, Field (Ai i)) + (_ : ∀ i, Algebra K (Ai i)) (_ : A ≃ₐ[K] Π i, Ai i), + ∀ i, Algebra.IsSeparable K (Ai i) := by + classical + constructor + · intro H + have := FormallyUnramified.finite_of_free K A + have := FormallyUnramified.isReduced_of_field K A + have : IsArtinianRing A := isArtinian_of_tower K inferInstance + letI : Fintype {I : Ideal A | I.IsMaximal} := (nonempty_fintype _).some + let v (i : {I : Ideal A | I.IsMaximal}) : A := (IsArtinianRing.equivPi A).symm (Pi.single i 1) + let e : A ≃ₐ[K] _ := { __ := IsArtinianRing.equivPi A, commutes' := fun r ↦ rfl } + have := (FormallyEtale.iff_of_equiv e).mp inferInstance + rw [FormallyEtale.pi_iff] at this + exact ⟨_, inferInstance, _, _, _, e, fun I ↦ (iff_isSeparable _ _).mp inferInstance⟩ + · intro ⟨I, _, Ai, _, _, e, _⟩ + rw [FormallyEtale.iff_of_equiv e, FormallyEtale.pi_iff] + have (i) : EssFiniteType K (Ai i) := by + letI := ((Pi.evalRingHom Ai i).comp e.toRingHom).toAlgebra + have : IsScalarTower K A (Ai i) := + .of_algebraMap_eq fun r ↦ by simp [RingHom.algebraMap_toAlgebra] + have : Algebra.FiniteType A (Ai i) := .of_surjective inferInstance (Algebra.ofId _ _) + (RingHomSurjective.is_surjective (σ := Pi.evalRingHom Ai i).comp e.surjective) + exact EssFiniteType.comp K A (Ai i) + exact fun I ↦ (iff_isSeparable _ _).mpr inferInstance + end Algebra.FormallyEtale + +attribute [local instance] + IsArtinianRing.subtype_isMaximal_finite IsArtinianRing.fieldOfSubtypeIsMaximal in +/-- +`A` is étale over a field `K` if and only if +`A` is a finite product of finite separable field extensions. +-/ +theorem Algebra.Etale.iff_exists_algEquiv_prod : + Etale K A ↔ + ∃ (I : Type u) (_ : Finite I) (Ai : I → Type u) (_ : ∀ i, Field (Ai i)) + (_ : ∀ i, Algebra K (Ai i)) (_ : A ≃ₐ[K] Π i, Ai i), + ∀ i, Module.Finite K (Ai i) ∧ Algebra.IsSeparable K (Ai i) := by + constructor + · intro H + obtain ⟨I, _, Ai, _, _, e, _⟩ := (FormallyEtale.iff_exists_algEquiv_prod K A).mp inferInstance + have := FormallyUnramified.finite_of_free K A + exact ⟨_, ‹_›, _, _, _, e, fun i ↦ ⟨.of_surjective ((LinearMap.proj i).comp e.toLinearMap) + ((Function.surjective_eval i).comp e.surjective), inferInstance⟩⟩ + · intro ⟨I, _, Ai, _, _, e, H⟩ + choose h₁ h₂ using H + have := Module.Finite.of_surjective e.symm.toLinearMap e.symm.surjective + refine ⟨?_, FinitePresentation.of_finiteType.mp inferInstance⟩ + exact (FormallyEtale.iff_exists_algEquiv_prod K A).mpr ⟨_, inferInstance, _, _, _, e, h₂⟩ diff --git a/Mathlib/RingTheory/Etale/Pi.lean b/Mathlib/RingTheory/Etale/Pi.lean new file mode 100644 index 0000000000000..0e50eccd56627 --- /dev/null +++ b/Mathlib/RingTheory/Etale/Pi.lean @@ -0,0 +1,36 @@ +/- +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.Smooth.Pi +import Mathlib.RingTheory.Unramified.Pi +import Mathlib.RingTheory.Etale.Basic + +/-! + +# Formal-étaleness of finite products of rings + +## Main result + +- `Algebra.FormallyEtale.pi_iff`: If `I` is finite, `Π i : I, A i` is `R`-formally-étale + if and only if each `A i` is `R`-formally-étale. + +-/ + +universe u v + +namespace Algebra.FormallyEtale + +variable {R : Type (max u v)} {I : Type u} (A : I → Type (max u v)) +variable [CommRing R] [∀ i, CommRing (A i)] [∀ i, Algebra R (A i)] + +theorem pi_iff [Finite I] : + FormallyEtale R (Π i, A i) ↔ ∀ i, FormallyEtale R (A i) := by + simp_rw [FormallyEtale.iff_unramified_and_smooth, forall_and] + rw [FormallyUnramified.pi_iff A, FormallySmooth.pi_iff A] + +instance [Finite I] [∀ i, FormallyEtale R (A i)] : FormallyEtale R (Π i, A i) := + .of_unramified_and_smooth + +end Algebra.FormallyEtale diff --git a/Mathlib/RingTheory/Smooth/Pi.lean b/Mathlib/RingTheory/Smooth/Pi.lean index cfb4cf1dd16df..aab45da1545e8 100644 --- a/Mathlib/RingTheory/Smooth/Pi.lean +++ b/Mathlib/RingTheory/Smooth/Pi.lean @@ -118,4 +118,7 @@ theorem pi_iff [Finite I] : rwa [mul_sub, ← map_mul, mul_comm, mul_assoc, (he.idem i).eq, he', ← map_mul, ← Pi.single_mul, one_mul, sub_eq_zero] at this +instance [Finite I] [∀ i, FormallySmooth R (A i)] : FormallySmooth R (Π i, A i) := + (pi_iff _).mpr ‹_› + end Algebra.FormallySmooth diff --git a/Mathlib/RingTheory/Unramified/Pi.lean b/Mathlib/RingTheory/Unramified/Pi.lean index 755c16a601f58..ddebc82f3785d 100644 --- a/Mathlib/RingTheory/Unramified/Pi.lean +++ b/Mathlib/RingTheory/Unramified/Pi.lean @@ -95,4 +95,7 @@ theorem pi_iff : Ideal.Quotient.mkₐ_eq_mk] exact Ideal.mem_map_of_mem (Ideal.Quotient.mk J') (hf (Pi.single x r)) +instance [∀ i, FormallyUnramified R (f i)] : FormallyUnramified R (Π i, f i) := + (pi_iff _).mpr ‹_› + end Algebra.FormallyUnramified From f0c7c5c091e494feb1f8abf7a8d0e71815f4b06a Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 6 Jan 2025 14:36:15 +0000 Subject: [PATCH 183/189] feat(ContinuousMultilinearMap): add lemmas about `.prod` (#20462) --- .../NormedSpace/Multilinear/Basic.lean | 7 +- .../Algebra/Module/Multilinear/Basic.lean | 67 +++++++++++++++++-- 2 files changed, 62 insertions(+), 12 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index e45477c7b3695..c6c4a0ad92ef2 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -610,14 +610,9 @@ variable (𝕜 E E' G G') def prodL : ContinuousMultilinearMap 𝕜 E G × ContinuousMultilinearMap 𝕜 E G' ≃ₗᵢ[𝕜] ContinuousMultilinearMap 𝕜 E (G × G') where - toFun f := f.1.prod f.2 - invFun f := - ((ContinuousLinearMap.fst 𝕜 G G').compContinuousMultilinearMap f, - (ContinuousLinearMap.snd 𝕜 G G').compContinuousMultilinearMap f) + __ := prodEquiv map_add' _ _ := rfl map_smul' _ _ := rfl - left_inv f := by ext <;> rfl - right_inv f := by ext <;> rfl norm_map' f := opNorm_prod f.1 f.2 /-- `ContinuousMultilinearMap.pi` as a `LinearIsometryEquiv`. -/ diff --git a/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean b/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean index 65348675488a9..442b6637f5b9f 100644 --- a/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean @@ -299,6 +299,55 @@ theorem _root_.ContinuousLinearMap.compContinuousMultilinearMap_coe (g : M₂ ext m rfl +/-- `ContinuousMultilinearMap.prod` as an `Equiv`. -/ +@[simps apply symm_apply_fst symm_apply_snd, simps (config := .lemmasOnly) symm_apply] +def prodEquiv : + (ContinuousMultilinearMap R M₁ M₂ × ContinuousMultilinearMap R M₁ M₃) ≃ + ContinuousMultilinearMap R M₁ (M₂ × M₃) where + toFun f := f.1.prod f.2 + invFun f := ((ContinuousLinearMap.fst _ _ _).compContinuousMultilinearMap f, + (ContinuousLinearMap.snd _ _ _).compContinuousMultilinearMap f) + left_inv _ := rfl + right_inv _ := rfl + +theorem prod_ext_iff {f g : ContinuousMultilinearMap R M₁ (M₂ × M₃)} : + f = g ↔ (ContinuousLinearMap.fst _ _ _).compContinuousMultilinearMap f = + (ContinuousLinearMap.fst _ _ _).compContinuousMultilinearMap g ∧ + (ContinuousLinearMap.snd _ _ _).compContinuousMultilinearMap f = + (ContinuousLinearMap.snd _ _ _).compContinuousMultilinearMap g := by + rw [← Prod.mk.inj_iff, ← prodEquiv_symm_apply, ← prodEquiv_symm_apply, Equiv.apply_eq_iff_eq] + +@[ext] +theorem prod_ext {f g : ContinuousMultilinearMap R M₁ (M₂ × M₃)} + (h₁ : (ContinuousLinearMap.fst _ _ _).compContinuousMultilinearMap f = + (ContinuousLinearMap.fst _ _ _).compContinuousMultilinearMap g) + (h₂ : (ContinuousLinearMap.snd _ _ _).compContinuousMultilinearMap f = + (ContinuousLinearMap.snd _ _ _).compContinuousMultilinearMap g) : f = g := + prod_ext_iff.mpr ⟨h₁, h₂⟩ + +theorem eq_prod_iff {f : ContinuousMultilinearMap R M₁ (M₂ × M₃)} + {g : ContinuousMultilinearMap R M₁ M₂} {h : ContinuousMultilinearMap R M₁ M₃} : + f = g.prod h ↔ (ContinuousLinearMap.fst _ _ _).compContinuousMultilinearMap f = g ∧ + (ContinuousLinearMap.snd _ _ _).compContinuousMultilinearMap f = h := + prod_ext_iff + +theorem add_prod_add [ContinuousAdd M₂] [ContinuousAdd M₃] + (f₁ f₂ : ContinuousMultilinearMap R M₁ M₂) (g₁ g₂ : ContinuousMultilinearMap R M₁ M₃) : + (f₁ + f₂).prod (g₁ + g₂) = f₁.prod g₁ + f₂.prod g₂ := + rfl + +theorem smul_prod_smul {S : Type*} [Monoid S] [DistribMulAction S M₂] [DistribMulAction S M₃] + [ContinuousConstSMul S M₂] [SMulCommClass R S M₂] + [ContinuousConstSMul S M₃] [SMulCommClass R S M₃] + (c : S) (f : ContinuousMultilinearMap R M₁ M₂) (g : ContinuousMultilinearMap R M₁ M₃) : + (c • f).prod (c • g) = c • f.prod g := + rfl + +@[simp] +theorem zero_prod_zero : + (0 : ContinuousMultilinearMap R M₁ M₂).prod (0 : ContinuousMultilinearMap R M₁ M₃) = 0 := + rfl + /-- `ContinuousMultilinearMap.pi` as an `Equiv`. -/ @[simps] def piEquiv {ι' : Type*} {M' : ι' → Type*} [∀ i, AddCommMonoid (M' i)] @@ -306,12 +355,8 @@ def piEquiv {ι' : Type*} {M' : ι' → Type*} [∀ i, AddCommMonoid (M' i)] (∀ i, ContinuousMultilinearMap R M₁ (M' i)) ≃ ContinuousMultilinearMap R M₁ (∀ i, M' i) where toFun := ContinuousMultilinearMap.pi invFun f i := (ContinuousLinearMap.proj i : _ →L[R] M' i).compContinuousMultilinearMap f - left_inv f := by - ext - rfl - right_inv f := by - ext - rfl + left_inv _ := rfl + right_inv _ := rfl /-- An equivalence of the index set defines an equivalence between the spaces of continuous multilinear maps. This is the forward map of this equivalence. -/ @@ -453,6 +498,16 @@ instance : AddCommGroup (ContinuousMultilinearMap R M₁ M₂) := toMultilinearMap_injective.addCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl +theorem neg_prod_neg [AddCommGroup M₃] [Module R M₃] [TopologicalSpace M₃] [TopologicalAddGroup M₃] + (f : ContinuousMultilinearMap R M₁ M₂) (g : ContinuousMultilinearMap R M₁ M₃) : + (-f).prod (-g) = - f.prod g := + rfl + +theorem sub_prod_sub [AddCommGroup M₃] [Module R M₃] [TopologicalSpace M₃] [TopologicalAddGroup M₃] + (f₁ f₂ : ContinuousMultilinearMap R M₁ M₂) (g₁ g₂ : ContinuousMultilinearMap R M₁ M₃) : + (f₁ - f₂).prod (g₁ - g₂) = f₁.prod g₁ - f₂.prod g₂ := + rfl + end TopologicalAddGroup end Ring From 15b2729d1e0f9d80a5d712fc3734db52d8c60069 Mon Sep 17 00:00:00 2001 From: Daniel Weber Date: Mon, 6 Jan 2025 14:48:38 +0000 Subject: [PATCH 184/189] feat(Combinatorics/SimpleGraph/Walk): add `penultimate` and `snd` (#16769) Here are some reasons for it: - We want to add `penultimate`, as `p.getVert (p.length - 1)` duplicates `p`, and `snd` should exist for symmetry. - In some cases it makes sense to consider the vertices of a walk without the first/last one, and then this is the walk version of `List.head`/`List.last`, not `l[1]`/`l[l.size-2]` - In walks the first/last node and adjacency are a lot more important than they usually are in lists, so it makes more sense to have a special term for the node adjacent to the first/last node. - Walks always have a sensible default, so there's no need for the "take a proof/default value/panic/return an option" duplication which lists have and so this doesn't require much API. Co-authored-by: Daniel Weber <55664973+Command-Master@users.noreply.github.com> --- .../Combinatorics/SimpleGraph/Acyclic.lean | 2 +- .../Connectivity/WalkCounting.lean | 2 +- Mathlib/Combinatorics/SimpleGraph/Path.lean | 9 ++- Mathlib/Combinatorics/SimpleGraph/Walk.lean | 80 ++++++++++++++----- 4 files changed, 68 insertions(+), 25 deletions(-) diff --git a/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean b/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean index 9782814a525cc..126e91256aa32 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean @@ -175,7 +175,7 @@ lemma IsTree.card_edgeFinset [Fintype V] [Fintype G.edgeSet] (hG : G.IsTree) : 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 (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 diff --git a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean index 01d7e6a4232f3..10adfe87aff11 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean @@ -63,7 +63,7 @@ Note that `u` and `v` may be the same. -/ @[simps] def walkLengthTwoEquivCommonNeighbors (u v : V) : {p : G.Walk u v // p.length = 2} ≃ G.commonNeighbors u v where - toFun p := ⟨p.val.getVert 1, match p with + toFun p := ⟨p.val.snd, match p with | ⟨.cons _ (.cons _ .nil), _⟩ => ⟨‹G.Adj u _›, ‹G.Adj _ v›.symm⟩⟩ invFun w := ⟨w.prop.1.toWalk.concat w.prop.2.symm, rfl⟩ left_inv | ⟨.cons _ (.cons _ .nil), hp⟩ => by rfl diff --git a/Mathlib/Combinatorics/SimpleGraph/Path.lean b/Mathlib/Combinatorics/SimpleGraph/Path.lean index 34fc605b7283d..e95e8857cb5e8 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Path.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Path.lean @@ -258,10 +258,11 @@ lemma isCycle_reverse {p : G.Walk u u} : p.reverse.IsCycle ↔ p.IsCycle where mp h := by simpa using h.reverse mpr := .reverse -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 +lemma IsPath.tail {p : G.Walk u v} (hp : p.IsPath) : p.tail.IsPath := by + cases p with + | nil => simp + | cons hadj p => + simp_all [Walk.isPath_def] /-! ### About paths -/ diff --git a/Mathlib/Combinatorics/SimpleGraph/Walk.lean b/Mathlib/Combinatorics/SimpleGraph/Walk.lean index ae8224b0ec59d..af48c233e9bcd 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Walk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Walk.lean @@ -155,6 +155,9 @@ def getVert {u v : V} : G.Walk u v → ℕ → V @[simp] theorem getVert_zero {u v} (w : G.Walk u v) : w.getVert 0 = u := by cases w <;> rfl +@[simp] +theorem getVert_nil (u : V) {i : ℕ} : (@nil _ G u).getVert i = u := rfl + theorem getVert_of_length_le {u v} (w : G.Walk u v) {i : ℕ} (hi : w.length ≤ i) : w.getVert i = v := by induction w generalizing i with @@ -177,11 +180,6 @@ 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) : (p.cons h).getVert (n + 1) = p.getVert n := rfl @@ -829,10 +827,6 @@ 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 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 @@ -840,11 +834,49 @@ def drop {u v : V} (p : G.Walk u v) (n : ℕ) : G.Walk (p.getVert n) v := | 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) : G.Walk (p.getVert 1) v := p.drop 1 +/-- The second vertex of a walk, or the only vertex in a nil walk. -/ +abbrev snd (p : G.Walk u v) : V := p.getVert 1 + +@[simp] lemma adj_snd {p : G.Walk v w} (hp : ¬ p.Nil) : + G.Adj v p.snd := by + simpa using adj_getVert_succ p (by simpa [not_nil_iff_lt_length] using hp : 0 < p.length) + +lemma snd_cons {u v w} (q : G.Walk v w) (hadj : G.Adj u v) : + (q.cons hadj).snd = v := by simp + +/-- The penultimate vertex of a walk, or the only vertex in a nil walk. -/ +abbrev penultimate (p : G.Walk u v) : V := p.getVert (p.length - 1) + +@[simp] +lemma penultimate_cons_nil (h : G.Adj u v) : (cons h nil).penultimate = u := rfl + +@[simp] +lemma penultimate_cons_cons {w'} (h : G.Adj u v) (h₂ : G.Adj v w) (p : G.Walk w w') : + (cons h (cons h₂ p)).penultimate = (cons h₂ p).penultimate := rfl + +lemma penultimate_cons_of_not_nil (h : G.Adj u v) (p : G.Walk v w) (hp : ¬ p.Nil) : + (cons h p).penultimate = p.penultimate := + p.notNilRec (by simp) hp h + +@[simp] +lemma penultimate_concat {t u v} (p : G.Walk u v) (h : G.Adj v t) : + (p.concat h).penultimate = v := by simp [penultimate, concat_eq_append, getVert_append] + +@[simp] +lemma adj_penultimate {p : G.Walk v w} (hp : ¬ p.Nil) : + G.Adj p.penultimate w := by + conv => rhs; rw [← getVert_length p] + rw [nil_iff_length_eq] at hp + convert adj_getVert_succ _ _ <;> omega + +/-- The walk obtained by removing the first dart of a walk. A nil walk stays nil. -/ +def tail (p : G.Walk u v) : G.Walk (p.snd) v := p.drop 1 + +@[simp] +lemma tail_nil : (@nil _ G v).tail = .nil := rfl @[simp] -lemma tail_cons_nil (h : G.Adj u v) : (Walk.cons h .nil).tail = .nil := by rfl +lemma tail_cons_nil (h : G.Adj u v) : (Walk.cons h .nil).tail = .nil := 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 @@ -856,16 +888,26 @@ lemma tail_cons_eq (h : G.Adj u v) (p : G.Walk v w) : @[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.snd + adj := p.adj_snd hp + +/-- The last dart of a walk. -/ +@[simps] +def lastDart (p : G.Walk v w) (hp : ¬ p.Nil) : G.Dart where + fst := p.penultimate + snd := w + adj := p.adj_penultimate 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.snd) := rfl + +lemma edge_lastDart (p : G.Walk v w) (hp : ¬ p.Nil) : + (p.lastDart hp).edge = s(p.penultimate, w) := 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 + cons (p.adj_snd hp) p.tail = p := by cases p with | nil => simp only [nil_nil, not_true_eq_false] at hp | cons h q => @@ -1085,7 +1127,7 @@ theorem exists_boundary_dart {u v : V} (p : G.Walk u v) (S : Set V) (uS : u ∈ | .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) : +@[simp] lemma getVert_tail {u v n} (p : G.Walk u v) : p.tail.getVert n = p.getVert (n + 1) := by match p with | .nil => rfl @@ -1121,8 +1163,8 @@ theorem mem_support_iff_exists_getVert {u v w : V} {p : G.Walk v w} : 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 + rw [getVert_tail, length_tail_add_one hnp] + have : (n - 1 + 1) = n := by omega rwa [this] termination_by p.length decreasing_by From 9ab795afa4a2d4098113c0e893a6a95969ea4cd2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 6 Jan 2025 15:20:42 +0000 Subject: [PATCH 185/189] feat(1000): fill in more entries (#20470) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- [![Open in Gitpod](https://gitpod.io/button/open-in-gitpod.svg)](https://gitpod.io/from-referrer/) --------- Co-authored-by: Rémy Degenne --- .../InnerProductSpace/Projection.lean | 6 +- Mathlib/Combinatorics/SetFamily/LYM.lean | 2 +- .../MeasureTheory/Covering/Besicovitch.lean | 27 ++-- Mathlib/MeasureTheory/Covering/Vitali.lean | 14 +- docs/1000.yaml | 142 +++++++++++++++++- 5 files changed, 167 insertions(+), 24 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/Projection.lean b/Mathlib/Analysis/InnerProductSpace/Projection.lean index 902821ff1cdfd..ca88319d17068 100644 --- a/Mathlib/Analysis/InnerProductSpace/Projection.lean +++ b/Mathlib/Analysis/InnerProductSpace/Projection.lean @@ -63,10 +63,10 @@ local notation "absR" => abs -- FIXME this monolithic proof causes a deterministic timeout with `-T50000` -- It should be broken in a sequence of more manageable pieces, -- perhaps with individual statements for the three steps below. -/-- Existence of minimizers +/-- **Existence of minimizers**, aka the **Hilbert projection theorem**. + Let `u` be a point in a real inner product space, and let `K` be a nonempty complete convex subset. -Then there exists a (unique) `v` in `K` that minimizes the distance `‖u - v‖` to `u`. - -/ +Then there exists a (unique) `v` in `K` that minimizes the distance `‖u - v‖` to `u`. -/ theorem exists_norm_eq_iInf_of_complete_convex {K : Set F} (ne : K.Nonempty) (h₁ : IsComplete K) (h₂ : Convex ℝ K) : ∀ u : F, ∃ v ∈ K, ‖u - v‖ = ⨅ w : K, ‖u - w‖ := fun u => by let δ := ⨅ w : K, ‖u - w‖ diff --git a/Mathlib/Combinatorics/SetFamily/LYM.lean b/Mathlib/Combinatorics/SetFamily/LYM.lean index 67a603abd4165..24ec53fa6da54 100644 --- a/Mathlib/Combinatorics/SetFamily/LYM.lean +++ b/Mathlib/Combinatorics/SetFamily/LYM.lean @@ -211,7 +211,7 @@ end LYM /-- **Sperner's theorem**. The size of an antichain in `Finset α` is bounded by the size of the maximal layer in `Finset α`. This precisely means that `Finset α` is a Sperner order. -/ -theorem IsAntichain.sperner [Fintype α] {𝒜 : Finset (Finset α)} +theorem _root_.IsAntichain.sperner [Fintype α] {𝒜 : Finset (Finset α)} (h𝒜 : IsAntichain (· ⊆ ·) (𝒜 : Set (Finset α))) : #𝒜 ≤ (Fintype.card α).choose (Fintype.card α / 2) := by classical diff --git a/Mathlib/MeasureTheory/Covering/Besicovitch.lean b/Mathlib/MeasureTheory/Covering/Besicovitch.lean index a21afb385c697..9f8b24ead46f2 100644 --- a/Mathlib/MeasureTheory/Covering/Besicovitch.lean +++ b/Mathlib/MeasureTheory/Covering/Besicovitch.lean @@ -791,14 +791,17 @@ theorem exists_disjoint_closedBall_covering_ae_of_finiteMeasure_aux (μ : Measur rw [← Nat.succ_eq_add_one, u_succ] exact (hF (u n) (Pu n)).1 -/-- The measurable Besicovitch covering theorem. Assume that, for any `x` in a set `s`, -one is given a set of admissible closed balls centered at `x`, with arbitrarily small radii. -Then there exists a disjoint covering of almost all `s` by admissible closed balls centered at some -points of `s`. -This version requires that the underlying measure is sigma-finite, and that the space has the +/-- The measurable **Besicovitch covering theorem**. + +Assume that, for any `x` in a set `s`, one is given a set of admissible closed balls centered at +`x`, with arbitrarily small radii. Then there exists a disjoint covering of almost all `s` by +admissible closed balls centered at some points of `s`. + +This version requires the underlying measure to be sigma-finite, and the space to have the Besicovitch covering property (which is satisfied for instance by normed real vector spaces). It expresses the conclusion in a slightly awkward form (with a subset of `α × ℝ`) coming from the proof technique. + For a version giving the conclusion in a nicer form, see `exists_disjoint_closedBall_covering_ae`. -/ theorem exists_disjoint_closedBall_covering_ae_aux (μ : Measure α) [SFinite μ] (f : α → Set ℝ) @@ -813,12 +816,14 @@ theorem exists_disjoint_closedBall_covering_ae_aux (μ : Measure α) [SFinite μ ⟨t, t_count, ts, tr, tν, tdisj⟩ exact ⟨t, t_count, ts, tr, hμν tν, tdisj⟩ -/-- The measurable Besicovitch covering theorem. Assume that, for any `x` in a set `s`, -one is given a set of admissible closed balls centered at `x`, with arbitrarily small radii. -Then there exists a disjoint covering of almost all `s` by admissible closed balls centered at some -points of `s`. We can even require that the radius at `x` is bounded by a given function `R x`. -(Take `R = 1` if you don't need this additional feature). -This version requires that the underlying measure is sigma-finite, and that the space has the +/-- The measurable **Besicovitch covering theorem**. + +Assume that, for any `x` in a set `s`, one is given a set of admissible closed balls centered at +`x`, with arbitrarily small radii. Then there exists a disjoint covering of almost all `s` by +admissible closed balls centered at some points of `s`. We can even require that the radius at `x` +is bounded by a given function `R x`. (Take `R = 1` if you don't need this additional feature). + +This version requires the underlying measure to be sigma-finite, and the space to have the Besicovitch covering property (which is satisfied for instance by normed real vector spaces). -/ theorem exists_disjoint_closedBall_covering_ae (μ : Measure α) [SFinite μ] (f : α → Set ℝ) diff --git a/Mathlib/MeasureTheory/Covering/Vitali.lean b/Mathlib/MeasureTheory/Covering/Vitali.lean index 416a61a4c48c3..5b2934b6db6d8 100644 --- a/Mathlib/MeasureTheory/Covering/Vitali.lean +++ b/Mathlib/MeasureTheory/Covering/Vitali.lean @@ -198,12 +198,14 @@ theorem exists_disjoint_subfamily_covering_enlargement_closedBall alias exists_disjoint_subfamily_covering_enlargment_closedBall := exists_disjoint_subfamily_covering_enlargement_closedBall -/-- The measurable Vitali covering theorem. Assume one is given a family `t` of closed sets with -nonempty interior, such that each `a ∈ t` is included in a ball `B (x, r)` and covers a definite -proportion of the ball `B (x, 3 r)` for a given measure `μ` (think of the situation where `μ` is -a doubling measure and `t` is a family of balls). Consider a (possibly non-measurable) set `s` -at which the family is fine, i.e., every point of `s` belongs to arbitrarily small elements of `t`. -Then one can extract from `t` a disjoint subfamily that covers almost all `s`. +/-- The measurable **Vitali covering theorem**. + +Assume one is given a family `t` of closed sets with nonempty interior, such that each `a ∈ t` is +included in a ball `B (x, r)` and covers a definite proportion of the ball `B (x, 3 r)` for a given +measure `μ` (think of the situation where `μ` is a doubling measure and `t` is a family of balls). +Consider a (possibly non-measurable) set `s` at which the family is fine, i.e., every point of `s` +belongs to arbitrarily small elements of `t`. Then one can extract from `t` a disjoint subfamily +that covers almost all `s`. For more flexibility, we give a statement with a parameterized family of sets. -/ diff --git a/docs/1000.yaml b/docs/1000.yaml index 50d0e3fb4af8f..b7f982e4112d6 100644 --- a/docs/1000.yaml +++ b/docs/1000.yaml @@ -39,6 +39,9 @@ Q32182: Q33481: title: Arrow's impossibility theorem + author: Benjamin Davidson, Andrew Souther + date: 2021 + url: https://github.com/asouther4/lean-social-choice/blob/master/src/arrows_theorem.lean Q98831: title: Multiplication theorem @@ -54,9 +57,13 @@ Q132427: Q132469: title: Fermat's Last Theorem + # See https://imperialcollegelondon.github.io/FLT/ Q137164: title: Besicovitch covering theorem + decl: Besicovitch.exists_disjoint_closedBall_covering_ae + author: Sébastien Gouëzel + date: 2021 Q154210: title: CPCTC @@ -72,6 +79,9 @@ Q174955: Q179208: title: Cayley's theorem + decl: Equiv.Perm.subgroupOfMulAction + author: Eric Wieser + date: 2021 Q179467: title: Fourier theorem @@ -88,12 +98,17 @@ Q180345: Q182505: title: Bayes' theorem decl: ProbabilityTheory.cond_eq_inv_mul_cond_mul + author: Rishikesh Vaishnav + date: 2022 Q184410: title: Four color theorem Q188295: title: Fermat's little theorem + decl: ZMod.pow_card_sub_one_eq_one + author: Aaron Anderson + date: 2020 Q188745: title: Classification of Platonic solids @@ -104,6 +119,7 @@ Q189136: # XXX: other branch found exists_ratio_deriv_eq_ratio_slope, # is about Cauchy's vs Lagrange's Mean Value Theorem, TODO decide! author: Yury G. Kudryashov + date: 2019 Q190026: title: Ford's theorem @@ -119,16 +135,25 @@ Q190556: title: De Moivre's theorem decl: Complex.cos_add_sin_mul_I_pow author: Abhimanyu Pallavi Sudhir + date: 2019 Q191693: title: Lebesgue's decomposition theorem + decl: MeasureTheory.Measure.haveLebesgueDecomposition_of_sigmaFinite + author: Kexing Ying + date: 2021 Q192760: title: Fundamental theorem of algebra + decl: Complex.isAlgClosed + author: Chris Hughes + date: 2019 Q193286: title: Rolle's theorem decl: exists_deriv_eq_zero + author: Yury G. Kudryashov + date: 2019 Q193878: title: Chinese remainder theorem @@ -159,6 +184,7 @@ Q203565: title: Solutions of a general cubic equation decl: Theorems100.cubic_eq_zero_iff author: Jeoff Lee + date: 2022 Q204884: title: Löb's theorem @@ -529,6 +555,9 @@ Q608294: Q609612: title: Knaster–Tarski theorem + decl: fixedPoints.completeLattice + author: Kenny Lau + date: 2018 Q612021: title: Gibbard–Satterthwaite theorem @@ -563,6 +592,7 @@ Q642620: title: Krein–Milman theorem decl: closure_convexHull_extremePoints author: Yaël Dillies + date: 2022 Q643513: title: Riesz–Fischer theorem @@ -655,15 +685,22 @@ Q716171: title: Erdős–Ginzburg–Ziv theorem decl: ZMod.erdos_ginzburg_ziv author: Yaël Dillies + date: 2023 Q718875: title: Erdős–Ko–Rado theorem + decl: Finset.erdos_ko_rado + author: Bhavik Mehta, Yaël Dillies + date: 2020 Q719966: title: Toda's theorem Q720469: title: Chevalley–Warning theorem + decl: char_dvd_card_solutions_of_sum_lt + author: Johan Commelin + date: 2019 Q721695: title: Szemerédi–Trotter theorem @@ -702,6 +739,9 @@ Q744440: Q748233: title: Sylvester–Gallai theorem + author: Bhavik Mehta + url: https://github.com/YaelDillies/LeanCamCombi/blob/6a6a670f324b2af82ae17f21f9d51ac0bc859f6f/LeanCamCombi/SylvesterChvatal.lean#L610 + date: 2022 Q751120: title: Thue–Siegel–Roth theorem @@ -709,6 +749,8 @@ Q751120: Q752375: title: Extreme value theorem decl: IsCompact.exists_isMinOn + author: Sébastien Gouëzel + date: 2018 Q755986: title: Atiyah–Bott fixed-point theorem @@ -720,6 +762,7 @@ Q756946: title: Lagrange's four-square theorem decl: Nat.sum_four_squares author: Chris Hughes + date: 2019 Q764287: title: Van der Waerden's theorem @@ -777,6 +820,9 @@ Q830513: Q834025: title: Cauchy integral theorem + decl: Complex.circleIntegral_div_sub_of_differentiable_on_off_countable + author: Yury Kudryashov + date: 2021 Q834211: title: Wallace–Bolyai–Gerwien theorem @@ -793,6 +839,9 @@ Q841893: Q842953: title: Lebesgue's density theorem + decl: Besicovitch.ae_tendsto_measure_inter_div + author: Oliver Nash + date: 2022 Q844612: title: Brun's theorem @@ -967,8 +1016,10 @@ Q944297: # Banach open mapping theorem/Banach-Schauder theorem in functional analysis is in mathlib (ContinuousLinearMap.isOpenMap) Q948664: - title: Kneser's theorem - # TODO: is this in mathlib already? + title: Kneser's addition theorem + author: Mantas Bakšys, Yaël Dillies + url: https://github.com/YaelDillies/LeanCamCombi/blob/master/LeanCamCombi/Kneser/Kneser.lean + date: 2022 Q951327: title: Pasch's theorem @@ -1166,6 +1217,7 @@ Q1097021: title: Minkowski's theorem decl: MeasureTheory.exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure author: Alex J. Best and Yaël Dillies + date: 2021 Q1103054: title: Lions–Lax–Milgram theorem @@ -1364,6 +1416,9 @@ Q1242398: Q1243340: title: Birkhoff–Von Neumann theorem + decl: doublyStochastic_eq_convexHull_permMatrix + author: Bhavik Mehta + date: 2024 Q1249069: title: Richardson's theorem @@ -1471,6 +1526,8 @@ Q1425529: Q1426292: title: Banach–Steinhaus theorem decl: banach_steinhaus + author: Jireh Loreaux + date: 2021 Q1434158: title: Fluctuation dissipation theorem @@ -1551,6 +1608,8 @@ Q1568612: Q1568811: title: Hahn decomposition theorem decl: MeasureTheory.SignedMeasure.exists_isCompl_positive_negative + author: Kexing Ying + date: 2021 Q1572474: title: Lindemann–Weierstrass theorem @@ -1582,6 +1641,9 @@ Q1632301: Q1632433: title: Helly's theorem + author: Vasily Nesterov + decl: Convex.helly_theorem + date: 2023 Q1637085: title: Poynting's theorem @@ -1597,6 +1659,9 @@ Q1683356: Q1687147: title: Sprague–Grundy theorem + decl: SetTheory.PGame.equiv_nim_grundyValue + author: Fox Thomson + date: 2020 Q1694565: title: Vinogradov's theorem @@ -1733,6 +1798,9 @@ Q2022775: Q2027347: title: Optional stopping theorem + decl: MeasureTheory.submartingale_iff_expected_stoppedValue_mono + author: Kexing Ying, Rémy Degenne + date: 2022 Q2028341: title: Ado's theorem @@ -1825,10 +1893,15 @@ Q2226774: Q2226786: title: Sperner's theorem + decl: IsAntichain.sperner + author: Bhavik Mehta, Alena Gusakov, Yaël Dillies + date: 2022 Q2226800: title: Schur–Zassenhaus theorem decl: Subgroup.exists_left_complement'_of_coprime + author: Thomas Browning + date: 2021 Q2226803: title: Tennenbaum's theorem @@ -1837,7 +1910,10 @@ Q2226807: title: Vantieghems theorem Q2226855: - title: Sarkovskii's theorem + title: Sharkovskii's theorem + # TODO: Find the code + author: Bhavik Mehta + date: 2022 Q2226868: title: Geometric mean theorem @@ -1864,6 +1940,9 @@ Q2270905: Q2275191: title: Lebesgue differentiation theorem + decl: VitaliFamily.ae_tendsto_lintegral_div + author: Sébastien Gouëzel + date: 2021 Q2277040: title: Shannon's expansion theorem @@ -1882,6 +1961,7 @@ Q2310718: Q2338929: title: Exchange theorem + # aka Steinitz exchange lemma. This must definitely be in mathlib somewhere Q2345282: title: Shannon's theorem @@ -1933,6 +2013,7 @@ Q2478371: Q2482753: title: Mann's theorem + # Done in LeanCamCombi Q2495664: title: Universal coefficient theorem @@ -1951,6 +2032,7 @@ Q2524990: Q2525646: title: Jordan–Hölder theorem + decl: CompositionSeries.jordan_holder Q2533936: title: Descartes's theorem on total angular defect @@ -2011,6 +2093,9 @@ Q2916568: Q2919401: title: Ostrowski's theorem + decl: Rat.AbsoluteValue.equiv_real_or_padic + author: 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, Francesco Veneziano + date: 2024 Q2981012: title: Monge's theorem @@ -2056,6 +2141,9 @@ Q3229335: Q3229352: title: Vitali covering theorem + decl: Vitali.exists_disjoint_covering_ae + author: Sébastien Gouëzel + date: 2021 Q3345678: title: Moore–Aronszajn theorem @@ -2092,6 +2180,9 @@ Q3526993: Q3526996: title: Kolmogorov extension theorem + author: Rémy Degenne, Peter Pfaffelhuber + date: 2023 + url: https://github.com/RemyDegenne/kolmogorov_extension4 Q3526998: title: Excision theorem @@ -2161,6 +2252,9 @@ Q3527100: Q3527102: title: Kruskal–Katona theorem + decl: Finset.kruskal_katona + author: Bhavik Mehta, Yaël Dillies + date: 2020 Q3527110: title: Lax–Wendroff theorem @@ -2212,6 +2306,9 @@ Q3527214: Q3527215: title: Hilbert projection theorem + decl: exists_norm_eq_iInf_of_complete_convex + author: Zhouhang Zhou + date: 2019 Q3527217: title: M. Riesz extension theorem @@ -2242,9 +2339,15 @@ Q3527247: Q3527250: title: Hadamard three-lines theorem + decl: Complex.HadamardThreeLines.norm_le_interpStrip_of_mem_verticalClosedStrip + author: Xavier Généreux + date: 2023 Q3527263: title: Kleene fixed-point theorem + decl: fixedPoints.lfp_eq_sSup_iterate + author: Ira Fesefeldt + date: 2024 Q3527279: title: Wiener's tauberian theorem @@ -2299,6 +2402,9 @@ Q3984052: Q3984053: title: Fourier inversion theorem + decl: Continuous.fourier_inversion + author: Sébastien Gouëzel + date: 2024 Q3984056: title: No-communication theorem @@ -2329,6 +2435,9 @@ Q4272645: Q4378868: title: Phragmén–Lindelöf theorem + decl: PhragmenLindelof.horizontal_strip + author: Yury G. Kudryashov + date: 2022 Q4378889: title: Carathéodory's theorem @@ -2416,6 +2525,9 @@ Q4677985: Q4695203: title: Four functions theorem + decl: four_functions_theorem + author: Yaël Dillies + date: 2023 Q4700718: title: Akhiezer's theorem @@ -2485,6 +2597,9 @@ Q4827308: Q4830725: title: Ax–Grothendieck theorem + decl: ax_grothendieck_zeroLocus + author: Chris Hughes + date: 2023 Q4832965: title: Aztec diamond theorem @@ -2698,6 +2813,9 @@ Q5166389: Q5171652: title: Corners theorem + decl: corners_theorem_nat + author: Yaël Dillies, Bhavik Mehta + date: 2022 Q5172143: title: Corona theorem @@ -3366,6 +3484,9 @@ Q7433034: Q7433182: title: Schwartz–Zippel theorem + decl: MvPolynomial.schwartz_zippel_sup_sum + author: Bolton Bailey, Yaël Dillies + date: 2023 Q7433295: title: Osterwalder–Schrader theorem @@ -3524,6 +3645,9 @@ Q7857350: Q7888360: title: Structure theorem for finitely generated modules over a principal ideal domain + decl: Module.equiv_free_prod_directSum + author: Pierre-Alexandre Bazin + date: 2022 Q7894110: title: Universal approximation theorem @@ -3587,6 +3711,9 @@ Q8066795: Q8074796: title: Zsigmondy's theorem + # TODO: Find the code + author: Mantas Bakšys + date: 2022 Q8081891: title: Śleszyński–Pringsheim theorem @@ -3605,6 +3732,9 @@ Q10942247: Q11352023: title: Vitali convergence theorem + decl: MeasureTheory.tendstoInMeasure_iff_tendsto_Lp + author: Igor Khavkine + date: 2024 Q11573495: title: Plancherel theorem for spherical functions @@ -3641,6 +3771,9 @@ Q15895894: Q16251580: title: Matiyasevich's theorem + decl: Pell.matiyasevic + author: Mario Carneiro + date: 2017 Q16680059: title: Strong perfect graph theorem @@ -3659,6 +3792,9 @@ Q17003552: Q17005116: title: Birkhoff's representation theorem + decl: LatticeHom.birkhoffSet + author: Yaël Dillies + date: 2022 Q17008559: title: Davenport–Schmidt theorem From 5716319cc527c44c374b07889bbd7c5c12965bdd Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Mon, 6 Jan 2025 15:23:00 +0000 Subject: [PATCH 186/189] fix: linkfix in 100.yaml (#20517) --- docs/100.yaml | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/docs/100.yaml b/docs/100.yaml index 991d6a09ae840..e7f8eb83b68c0 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -20,9 +20,8 @@ title : Gödel’s Incompleteness Theorem author : Shogo Saito links : - results : - - First: https://github.com/FormalizedFormalLogic/Incompleteness/blob/master/Incompleteness/Arith/First.lean - - Second: https://github.com/FormalizedFormalLogic/Incompleteness/blob/master/Incompleteness/Arith/Second.lean + First incompleteness theorem: https://github.com/FormalizedFormalLogic/Incompleteness/blob/master/Incompleteness/Arith/First.lean + Second incompleteness theorem: https://github.com/FormalizedFormalLogic/Incompleteness/blob/master/Incompleteness/Arith/Second.lean website: https://formalizedformallogic.github.io/Book/ 7: title : Law of Quadratic Reciprocity From 648ca59824034ebd79298ec7250f99c7eb71f3b0 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 6 Jan 2025 15:56:35 +0000 Subject: [PATCH 187/189] feat: A disjoint union is finite iff all sets are finite, and all but finitely many are empty (#20457) From the Carleson project --- Mathlib/Data/Set/Finite/Lattice.lean | 29 ++++++++++++++++++++++++++++ 1 file changed, 29 insertions(+) diff --git a/Mathlib/Data/Set/Finite/Lattice.lean b/Mathlib/Data/Set/Finite/Lattice.lean index 85c9d8dbf4beb..0f87a472988fd 100644 --- a/Mathlib/Data/Set/Finite/Lattice.lean +++ b/Mathlib/Data/Set/Finite/Lattice.lean @@ -152,6 +152,35 @@ theorem Finite.iUnion {ι : Type*} {s : ι → Set α} {t : Set ι} (ht : t.Fini · rw [he i hi, mem_empty_iff_false] at hx contradiction +/-- An indexed union of pairwise disjoint sets is finite iff all sets are finite, and all but +finitely many are empty. -/ +lemma finite_iUnion_iff {ι : Type*} {s : ι → Set α} (hs : Pairwise fun i j ↦ Disjoint (s i) (s j)) : + (⋃ i, s i).Finite ↔ (∀ i, (s i).Finite) ∧ {i | (s i).Nonempty}.Finite where + mp h := by + refine ⟨fun i ↦ h.subset <| subset_iUnion _ _, ?_⟩ + let u (i : {i | (s i).Nonempty}) : ⋃ i, s i := ⟨i.2.choose, mem_iUnion.2 ⟨i.1, i.2.choose_spec⟩⟩ + have u_inj : Function.Injective u := by + rintro ⟨i, hi⟩ ⟨j, hj⟩ hij + ext + refine hs.eq <| not_disjoint_iff.2 ⟨u ⟨i, hi⟩, hi.choose_spec, ?_⟩ + rw [hij] + exact hj.choose_spec + have : Finite (⋃ i, s i) := h + exact .of_injective u u_inj + mpr h := h.2.iUnion (fun _ _ ↦ h.1 _) (by simp [not_nonempty_iff_eq_empty]) + +@[simp] lemma finite_iUnion_of_subsingleton {ι : Sort*} [Subsingleton ι] {s : ι → Set α} : + (⋃ i, s i).Finite ↔ ∀ i, (s i).Finite := by + rw [← iUnion_plift_down, finite_iUnion_iff _root_.Subsingleton.pairwise] + simp [PLift.forall, Finite.of_subsingleton] + +/-- An indexed union of pairwise disjoint sets is finite iff all sets are finite, and all but +finitely many are empty. -/ +lemma PairwiseDisjoint.finite_biUnion_iff {f : β → Set α} {s : Set β} (hs : s.PairwiseDisjoint f) : + (⋃ i ∈ s, f i).Finite ↔ (∀ i ∈ s, (f i).Finite) ∧ {i ∈ s | (f i).Nonempty}.Finite := by + rw [finite_iUnion_iff (by aesop (add unfold safe [Pairwise, PairwiseDisjoint, Set.Pairwise]))] + simp + section preimage variable {f : α → β} {s : Set β} From ee71c31a5c26f79ff825829748dec28897a08375 Mon Sep 17 00:00:00 2001 From: Seewoo Lee <49933279+seewoo5@users.noreply.github.com> Date: Mon, 6 Jan 2025 16:31:30 +0000 Subject: [PATCH 188/189] feat: Polynomial FLT (#18882) Prove the polynomial FLT, using Mason-Stothers theorem #15706. More generally, prove non-solvability of Fermat-Catalan equation: $ux^p + vx^q + wz^r = 0$ where $u, v, w \in k^\times$ are units and $p, q, r \ge 3$ with $pq + qr + rp \le pqr$. Also derive the `IsCoprime b c` and `IsCoprime c a` assumptions from the `IsCoprime a b` and `a + b + c = 0` ones in `Polynomial.abc`. Co-authored-by: Jineon Baek Co-authored-by: Jineon Baek <34874130+jcpaik@users.noreply.github.com> Co-authored-by: Seewoo Lee --- Mathlib.lean | 1 + Mathlib/NumberTheory/FLT/MasonStothers.lean | 19 +- Mathlib/NumberTheory/FLT/Polynomial.lean | 275 ++++++++++++++++++++ 3 files changed, 290 insertions(+), 5 deletions(-) create mode 100644 Mathlib/NumberTheory/FLT/Polynomial.lean diff --git a/Mathlib.lean b/Mathlib.lean index 2a2fd484ec0df..2bdbb9627ace3 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3894,6 +3894,7 @@ import Mathlib.NumberTheory.EulerProduct.ExpLog import Mathlib.NumberTheory.FLT.Basic import Mathlib.NumberTheory.FLT.Four import Mathlib.NumberTheory.FLT.MasonStothers +import Mathlib.NumberTheory.FLT.Polynomial import Mathlib.NumberTheory.FLT.Three import Mathlib.NumberTheory.FactorisationProperties import Mathlib.NumberTheory.Fermat diff --git a/Mathlib/NumberTheory/FLT/MasonStothers.lean b/Mathlib/NumberTheory/FLT/MasonStothers.lean index eee4ff81ac138..bd5d8a2c3a537 100644 --- a/Mathlib/NumberTheory/FLT/MasonStothers.lean +++ b/Mathlib/NumberTheory/FLT/MasonStothers.lean @@ -16,9 +16,6 @@ Proof is based on this online note by Franz Lemmermeyer http://www.fen.bilkent.e which is essentially based on Noah Snyder's paper "An Alternative Proof of Mason's Theorem", but slightly different. -## TODO - -Prove polynomial FLT using Mason-Stothers theorem. -/ open Polynomial UniqueFactorizationMonoid UniqueFactorizationDomain EuclideanDomain @@ -51,13 +48,25 @@ private theorem abc_subcall {a b c w : k[X]} {hw : w ≠ 0} (wab : w = wronskian exact Nat.add_lt_add_right abc_dr_ndeg_lt _ /-- **Polynomial ABC theorem.** -/ -theorem Polynomial.abc {a b c : k[X]} (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) (hab : IsCoprime a b) - (hbc : IsCoprime b c) (hca : IsCoprime c a) (hsum : a + b + c = 0) : +protected theorem Polynomial.abc + {a b c : k[X]} (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) + (hab : IsCoprime a b) (hsum : a + b + c = 0) : ( natDegree a + 1 ≤ (radical (a * b * c)).natDegree ∧ natDegree b + 1 ≤ (radical (a * b * c)).natDegree ∧ natDegree c + 1 ≤ (radical (a * b * c)).natDegree ) ∨ derivative a = 0 ∧ derivative b = 0 ∧ derivative c = 0 := by set w := wronskian a b with wab + have hbc : IsCoprime b c := by + rw [add_eq_zero_iff_neg_eq] at hsum + rw [← hsum, IsCoprime.neg_right_iff] + convert IsCoprime.add_mul_left_right hab.symm 1 + rw [mul_one] + have hsum' : b + c + a = 0 := by rwa [add_rotate] at hsum + have hca : IsCoprime c a := by + rw [add_eq_zero_iff_neg_eq] at hsum' + rw [← hsum', IsCoprime.neg_right_iff] + convert IsCoprime.add_mul_left_right hbc.symm 1 + rw [mul_one] have wbc : w = wronskian b c := wronskian_eq_of_sum_zero hsum have wca : w = wronskian c a := by rw [add_rotate] at hsum diff --git a/Mathlib/NumberTheory/FLT/Polynomial.lean b/Mathlib/NumberTheory/FLT/Polynomial.lean new file mode 100644 index 0000000000000..e6beb32165d5a --- /dev/null +++ b/Mathlib/NumberTheory/FLT/Polynomial.lean @@ -0,0 +1,275 @@ +/- +Copyright (c) 2024 Jineon Baek 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.Algebra.Polynomial.Expand +import Mathlib.Algebra.GroupWithZero.Defs +import Mathlib.NumberTheory.FLT.Basic +import Mathlib.NumberTheory.FLT.MasonStothers +import Mathlib.RingTheory.Polynomial.Content + +/-! +# Fermat's Last Theorem for polynomials over a field + +This file states and proves the Fermat's Last Theorem for polynomials over a field. +For `n ≥ 3` not divisible by the characteristic of the coefficient field `k` and (pairwise) nonzero +coprime polynomials `a, b, c` (over a field) with `a ^ n + b ^ n = c ^ n`, +all polynomials must be constants. + +More generally, we can prove non-solvability of the Fermat-Catalan equation: there are no +non-constant polynomial solutions to the equation `u * a ^ p + v * b ^ q + w * c ^ r = 0`, where +`p, q, r ≥ 3` with `p * q + q * r + r * p ≤ p * q * r` , `p, q, r` not divisible by `char k`, +and `u, v, w` are nonzero elements in `k`. +FLT is the special case where `p = q = r = n`, `u = v = 1`, and `w = -1`. + +The proof uses the Mason-Stothers theorem (Polynomial ABC theorem) and infinite descent +(in the characteristic p case). +-/ + +open Polynomial UniqueFactorizationMonoid UniqueFactorizationDomain + +variable {k R : Type*} [Field k] [CommRing R] [IsDomain R] [NormalizationMonoid R] + [UniqueFactorizationMonoid R] + +private lemma Ne.isUnit_C {u : k} (hu : u ≠ 0) : IsUnit (C u) := + Polynomial.isUnit_C.mpr hu.isUnit + +-- auxiliary lemma that 'rotates' coprimality +private lemma rot_coprime + {p q r : ℕ} {a b c : k[X]} {u v w : k} + {hp : 0 < p} {hq : 0 < q} {hr : 0 < r} + {hu : u ≠ 0} {hv : v ≠ 0} {hw : w ≠ 0} + (heq : C u * a ^ p + C v * b ^ q + C w * c ^ r = 0) (hab : IsCoprime a b) : IsCoprime b c := by + have hCu : IsUnit (C u) := Ne.isUnit_C hu + have hCv : IsUnit (C v) := Ne.isUnit_C hv + have hCw : IsUnit (C w) := Ne.isUnit_C hw + rw [← IsCoprime.pow_iff hp hq, ← isCoprime_mul_units_left hCu hCv] at hab + rw [add_eq_zero_iff_neg_eq] at heq + rw [← IsCoprime.pow_iff hq hr, ← isCoprime_mul_units_left hCv hCw, + ← heq, IsCoprime.neg_right_iff] + convert IsCoprime.add_mul_left_right hab.symm 1 using 2 + rw [mul_one] + +private lemma ineq_pqr_contradiction {p q r a b c : ℕ} + (hp : 0 < p) (hq : 0 < q) (hr : 0 < r) + (hineq : q * r + r * p + p * q ≤ p * q * r) + (hpa : p * a < a + b + c) + (hqb : q * b < a + b + c) + (hrc : r * c < a + b + c) : False := by + suffices h : p * q * r * (a + b + c) + 1 ≤ p * q * r * (a + b + c) by simp at h + calc + _ = (p * a) * (q * r) + (q * b) * (r * p) + (r * c) * (p * q) + 1 := by ring + _ ≤ (a + b + c) * (q * r) + (a + b + c) * (r * p) + (a + b + c) * (p * q) := by + rw [Nat.succ_le] + gcongr + _ = (q * r + r * p + p * q) * (a + b + c) := by ring + _ ≤ _ := by gcongr + +private theorem Polynomial.flt_catalan_deriv [DecidableEq k] + {p q r : ℕ} (hp : 0 < p) (hq : 0 < q) (hr : 0 < r) + (hineq : q * r + r * p + p * q ≤ p * q * r) + (chp : (p : k) ≠ 0) (chq : (q : k) ≠ 0) (chr : (r : k) ≠ 0) + {a b c : k[X]} (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) + (hab : IsCoprime a b) {u v w : k} (hu : u ≠ 0) (hv : v ≠ 0) (hw : w ≠ 0) + (heq : C u * a ^ p + C v * b ^ q + C w * c ^ r = 0) : + derivative a = 0 ∧ derivative b = 0 ∧ derivative c = 0 := by + have hbc : IsCoprime b c := by apply rot_coprime heq <;> assumption + have hca : IsCoprime c a := by + rw [add_rotate] at heq; apply rot_coprime heq <;> assumption + have hCu := C_ne_zero.mpr hu + have hCv := C_ne_zero.mpr hv + have hCw := C_ne_zero.mpr hw + have hap := pow_ne_zero p ha + have hbq := pow_ne_zero q hb + have hcr := pow_ne_zero r hc + have habp : IsCoprime (C u * a ^ p) (C v * b ^ q) := by + rw [isCoprime_mul_units_left hu.isUnit_C hv.isUnit_C]; exact hab.pow + have hbcp : IsCoprime (C v * b ^ q) (C w * c ^ r) := by + rw [isCoprime_mul_units_left hv.isUnit_C hw.isUnit_C]; exact hbc.pow + have hcap : IsCoprime (C w * c ^ r) (C u * a ^ p) := by + rw [isCoprime_mul_units_left hw.isUnit_C hu.isUnit_C]; exact hca.pow + have habcp := hcap.symm.mul_left hbcp + + -- Use Mason-Stothers theorem + rcases Polynomial.abc + (mul_ne_zero hCu hap) (mul_ne_zero hCv hbq) (mul_ne_zero hCw hcr) + habp heq with nd_lt | dr0 + · simp_rw [radical_mul habcp, radical_mul habp, + radical_mul_of_isUnit_left hu.isUnit_C, + radical_mul_of_isUnit_left hv.isUnit_C, + radical_mul_of_isUnit_left hw.isUnit_C, + radical_pow a hp, radical_pow b hq, radical_pow c hr, + natDegree_mul hCu hap, + natDegree_mul hCv hbq, + natDegree_mul hCw hcr, + natDegree_C, natDegree_pow, zero_add, + ← radical_mul hab, + ← radical_mul (hca.symm.mul_left hbc)] at nd_lt + + obtain ⟨hpa', hqb', hrc'⟩ := nd_lt + have hpa := hpa'.trans natDegree_radical_le + have hqb := hqb'.trans natDegree_radical_le + have hrc := hrc'.trans natDegree_radical_le + rw [natDegree_mul (mul_ne_zero ha hb) hc, + natDegree_mul ha hb, Nat.add_one_le_iff] at hpa hqb hrc + exfalso + exact (ineq_pqr_contradiction hp hq hr hineq hpa hqb hrc) + · rw [derivative_C_mul, derivative_C_mul, derivative_C_mul, + mul_eq_zero_iff_left (C_ne_zero.mpr hu), + mul_eq_zero_iff_left (C_ne_zero.mpr hv), + mul_eq_zero_iff_left (C_ne_zero.mpr hw), + derivative_pow_eq_zero chp, + derivative_pow_eq_zero chq, + derivative_pow_eq_zero chr] at dr0 + exact dr0 + +-- helper lemma that gives a baggage of small facts on `contract (ringChar k) a` +private lemma find_contract {a : k[X]} + (ha : a ≠ 0) (hda : derivative a = 0) (chn0 : ringChar k ≠ 0) : + ∃ ca, ca ≠ 0 ∧ + a = expand k (ringChar k) ca ∧ + a.natDegree = ca.natDegree * ringChar k := by + have heq := (expand_contract (ringChar k) hda chn0).symm + refine ⟨contract (ringChar k) a, ?_, heq, ?_⟩ + · intro h + rw [h, map_zero] at heq + exact ha heq + · rw [← natDegree_expand, ← heq] + +variable [DecidableEq k] + +private theorem Polynomial.flt_catalan_aux + {p q r : ℕ} {a b c : k[X]} {u v w : k} + (heq : C u * a ^ p + C v * b ^ q + C w * c ^ r = 0) + (hp : 0 < p) (hq : 0 < q) (hr : 0 < r) + (hineq : q * r + r * p + p * q ≤ p * q * r) + (chp : (p : k) ≠ 0) (chq : (q : k) ≠ 0) (chr : (r : k) ≠ 0) + (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) (hab : IsCoprime a b) + (hu : u ≠ 0) (hv : v ≠ 0) (hw : w ≠ 0) : + a.natDegree = 0 := by + cases' eq_or_ne (ringChar k) 0 with ch0 chn0 + -- characteristic zero + · obtain ⟨da, -, -⟩ := flt_catalan_deriv + hp hq hr hineq chp chq chr ha hb hc hab hu hv hw heq + have czk : CharZero k := by + apply charZero_of_inj_zero + intro n + rw [ringChar.spec, ch0] + exact zero_dvd_iff.mp + rw [eq_C_of_derivative_eq_zero da] + exact natDegree_C _ + -- characteristic p > 0 + · generalize eq_d : a.natDegree = d + -- set up infinite descent + -- strong induct on `d := a.natDegree` + induction d + using Nat.case_strong_induction_on + generalizing a b c ha hb hc hab heq with + | hz => rfl + | hi d ih_d => -- have derivatives of `a, b, c` zero + obtain ⟨ad, bd, cd⟩ := flt_catalan_deriv + hp hq hr hineq chp chq chr ha hb hc hab hu hv hw heq + -- find contracts `ca, cb, cc` so that `a(k) = ca(k^ch)` + obtain ⟨ca, ca_nz, eq_a, eq_deg_a⟩ := find_contract ha ad chn0 + obtain ⟨cb, cb_nz, eq_b, eq_deg_b⟩ := find_contract hb bd chn0 + obtain ⟨cc, cc_nz, eq_c, eq_deg_c⟩ := find_contract hc cd chn0 + set ch := ringChar k + suffices hca : ca.natDegree = 0 by + rw [← eq_d, eq_deg_a, hca, zero_mul] + by_contra hnca; apply hnca + apply ih_d _ _ _ ca_nz cb_nz cc_nz <;> clear ih_d <;> try rfl + · apply (isCoprime_expand chn0).mp + rwa [← eq_a, ← eq_b] + · have _ : ch ≠ 1 := CharP.ringChar_ne_one + have hch2 : 2 ≤ ch := by omega + rw [← add_le_add_iff_right 1, ← eq_d, eq_deg_a] + refine le_trans ?_ (Nat.mul_le_mul_left _ hch2) + omega + · rw [eq_a, eq_b, eq_c, ← expand_C ch u, ← expand_C ch v, ← expand_C ch w] at heq + simp_rw [← map_pow, ← map_mul, ← map_add] at heq + rwa [Polynomial.expand_eq_zero (zero_lt_iff.mpr chn0)] at heq + +/-- Nonsolvability of the Fermat-Catalan equation. -/ +theorem Polynomial.flt_catalan + {p q r : ℕ} (hp : 0 < p) (hq : 0 < q) (hr : 0 < r) + (hineq : q * r + r * p + p * q ≤ p * q * r) + (chp : (p : k) ≠ 0) (chq : (q : k) ≠ 0) (chr : (r : k) ≠ 0) + {a b c : k[X]} (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) (hab : IsCoprime a b) + {u v w : k} (hu : u ≠ 0) (hv : v ≠ 0) (hw : w ≠ 0) + (heq : C u * a ^ p + C v * b ^ q + C w * c ^ r = 0) : + a.natDegree = 0 ∧ b.natDegree = 0 ∧ c.natDegree = 0 := by + -- WLOG argument: essentially three times flt_catalan_aux + have hbc : IsCoprime b c := by + apply rot_coprime heq hab <;> assumption + have heq' : C v * b ^ q + C w * c ^ r + C u * a ^ p = 0 := by + rwa [add_rotate] at heq + have hca : IsCoprime c a := by + apply rot_coprime heq' hbc <;> assumption + refine ⟨?_, ?_, ?_⟩ + · apply flt_catalan_aux heq <;> assumption + · rw [add_rotate] at heq hineq + rw [mul_rotate] at hineq + apply flt_catalan_aux heq <;> assumption + · rw [← add_rotate] at heq hineq + rw [← mul_rotate] at hineq + apply flt_catalan_aux heq <;> assumption + +theorem Polynomial.flt + {n : ℕ} (hn : 3 ≤ n) (chn : (n : k) ≠ 0) + {a b c : k[X]} (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) + (hab : IsCoprime a b) (heq : a ^ n + b ^ n = c ^ n) : + a.natDegree = 0 ∧ b.natDegree = 0 ∧ c.natDegree = 0 := by + have hn' : 0 < n := by linarith + rw [← sub_eq_zero, ← one_mul (a ^ n), ← one_mul (b ^ n), ← one_mul (c ^ n), sub_eq_add_neg, + ← neg_mul] at heq + have hone : (1 : k[X]) = C 1 := by rfl + have hneg_one : (-1 : k[X]) = C (-1) := by simp only [map_neg, map_one] + simp_rw [hneg_one, hone] at heq + apply flt_catalan hn' hn' hn' _ + chn chn chn ha hb hc hab one_ne_zero one_ne_zero (neg_ne_zero.mpr one_ne_zero) heq + have eq_lhs : n * n + n * n + n * n = 3 * n * n := by ring + rw [eq_lhs, mul_assoc, mul_assoc] + exact Nat.mul_le_mul_right (n * n) hn + +theorem fermatLastTheoremWith'_polynomial {n : ℕ} (hn : 3 ≤ n) (chn : (n : k) ≠ 0) : + FermatLastTheoremWith' k[X] n := by + rw [FermatLastTheoremWith'] + intros a b c ha hb hc heq + obtain ⟨a', eq_a⟩ := gcd_dvd_left a b + obtain ⟨b', eq_b⟩ := gcd_dvd_right a b + set d := gcd a b + have hd : d ≠ 0 := gcd_ne_zero_of_left ha + rw [eq_a, eq_b, mul_pow, mul_pow, ← mul_add] at heq + have hdc : d ∣ c := by + have hn : 0 < n := by omega + have hdncn : d ^ n ∣ c ^ n := ⟨_, heq.symm⟩ + + rw [dvd_iff_normalizedFactors_le_normalizedFactors hd hc] + rw [dvd_iff_normalizedFactors_le_normalizedFactors + (pow_ne_zero n hd) (pow_ne_zero n hc), + normalizedFactors_pow, normalizedFactors_pow] at hdncn + simp_rw [Multiset.le_iff_count, Multiset.count_nsmul, + mul_le_mul_left hn] at hdncn ⊢ + exact hdncn + obtain ⟨c', eq_c⟩ := hdc + rw [eq_a, mul_ne_zero_iff] at ha + rw [eq_b, mul_ne_zero_iff] at hb + rw [eq_c, mul_ne_zero_iff] at hc + rw [mul_comm] at eq_a eq_b eq_c + refine ⟨d, a', b', c', ⟨eq_a, eq_b, eq_c⟩, ?_⟩ + rw [eq_c, mul_pow, mul_comm, mul_left_inj' (pow_ne_zero n hd)] at heq + suffices goal : a'.natDegree = 0 ∧ b'.natDegree = 0 ∧ c'.natDegree = 0 by + simp [natDegree_eq_zero] at goal + obtain ⟨⟨ca', ha'⟩, ⟨cb', hb'⟩, ⟨cc', hc'⟩⟩ := goal + rw [← ha', ← hb', ← hc'] + rw [← ha', C_ne_zero] at ha + rw [← hb', C_ne_zero] at hb + rw [← hc', C_ne_zero] at hc + exact ⟨ha.right.isUnit_C, hb.right.isUnit_C, hc.right.isUnit_C⟩ + apply flt hn chn ha.right hb.right hc.right _ heq + convert isCoprime_div_gcd_div_gcd _ + · exact EuclideanDomain.eq_div_of_mul_eq_left ha.left eq_a.symm + · exact EuclideanDomain.eq_div_of_mul_eq_left ha.left eq_b.symm + · rw [eq_b] + exact mul_ne_zero hb.right hb.left From 69241484e921357756c37b487c438dc7ba2250e1 Mon Sep 17 00:00:00 2001 From: Artie Khovanov Date: Mon, 6 Jan 2025 17:01:52 +0000 Subject: [PATCH 189/189] chore(Algebra): make more names consistent (#20449) * Follows on from #20272 * Rename some substructure constructions to align with `oneLE` / `nonneg` * Clarify some documentation Moves: * squareIn -> square * evenIn -> even * sumSqIn -> sumSq --- Mathlib/Algebra/Group/Subgroup/Even.lean | 48 ++++++++++++------------ Mathlib/Algebra/Ring/SumsOfSquares.lean | 22 ++++++----- 2 files changed, 36 insertions(+), 34 deletions(-) diff --git a/Mathlib/Algebra/Group/Subgroup/Even.lean b/Mathlib/Algebra/Group/Subgroup/Even.lean index b4735654ebe05..488a9b76c3b59 100644 --- a/Mathlib/Algebra/Group/Subgroup/Even.lean +++ b/Mathlib/Algebra/Group/Subgroup/Even.lean @@ -13,73 +13,73 @@ This file defines the subgroup of squares / even elements in an abelian group. -/ namespace Subsemigroup -variable {S : Type*} [CommSemigroup S] {a : S} +variable {S : Type*} [CommSemigroup S] variable (S) in /-- -In a commutative semigroup `S`, `Subsemigroup.squareIn S` is the subsemigroup of squares in `S`. +In a commutative semigroup `S`, `Subsemigroup.square S` is the subsemigroup of squares in `S`. -/ @[to_additive -"In a commutative additive semigroup `S`, the type `AddSubsemigroup.evenIn S` -is the subsemigroup of even elements of `S`."] -def squareIn : Subsemigroup S where +"In a commutative additive semigroup `S`, `AddSubsemigroup.even S` +is the subsemigroup of even elements in `S`."] +def square : Subsemigroup S where carrier := {s : S | IsSquare s} mul_mem' := IsSquare.mul @[to_additive (attr := simp)] -theorem mem_squareIn : a ∈ squareIn S ↔ IsSquare a := Iff.rfl +theorem mem_square {a : S} : a ∈ square S ↔ IsSquare a := Iff.rfl @[to_additive (attr := simp, norm_cast)] -theorem coe_squareIn : squareIn S = {s : S | IsSquare s} := rfl +theorem coe_square : square S = {s : S | IsSquare s} := rfl end Subsemigroup namespace Submonoid -variable {M : Type*} [CommMonoid M] {a : M} +variable {M : Type*} [CommMonoid M] variable (M) in /-- -In a commutative monoid `M`, `Submonoid.squareIn M` is the submonoid of squares in `M`. +In a commutative monoid `M`, `Submonoid.square M` is the submonoid of squares in `M`. -/ @[to_additive -"In a commutative additive monoid `M`, the type `AddSubmonoid.evenIn M` -is the submonoid of even elements of `M`."] -def squareIn : Submonoid M where - __ := Subsemigroup.squareIn M +"In a commutative additive monoid `M`, `AddSubmonoid.even M` +is the submonoid of even elements in `M`."] +def square : Submonoid M where + __ := Subsemigroup.square M one_mem' := IsSquare.one @[to_additive (attr := simp)] -theorem squareIn_toSubsemigroup : (squareIn M).toSubsemigroup = .squareIn M := rfl +theorem square_toSubsemigroup : (square M).toSubsemigroup = .square M := rfl @[to_additive (attr := simp)] -theorem mem_squareIn : a ∈ squareIn M ↔ IsSquare a := Iff.rfl +theorem mem_square {a : M} : a ∈ square M ↔ IsSquare a := Iff.rfl @[to_additive (attr := simp, norm_cast)] -theorem coe_squareIn : squareIn M = {s : M | IsSquare s} := rfl +theorem coe_square : square M = {s : M | IsSquare s} := rfl end Submonoid namespace Subgroup -variable {G : Type*} [CommGroup G] {a : G} +variable {G : Type*} [CommGroup G] variable (G) in /-- -In an abelian group `G`, `Subgroup.squareIn G` is the subgroup of squares in `G`. +In an abelian group `G`, `Subgroup.square G` is the subgroup of squares in `G`. -/ @[to_additive -"In an abelian additive group `G`, the type `AddSubgroup.evenIn G` is +"In an abelian additive group `G`, `AddSubgroup.even G` is the subgroup of even elements in `G`."] -def squareIn : Subgroup G where - __ := Submonoid.squareIn G +def square : Subgroup G where + __ := Submonoid.square G inv_mem' := IsSquare.inv @[to_additive (attr := simp)] -theorem squareIn_toSubmonoid : (squareIn G).toSubmonoid = .squareIn G := rfl +theorem square_toSubmonoid : (square G).toSubmonoid = .square G := rfl @[to_additive (attr := simp)] -theorem mem_squareIn : a ∈ squareIn G ↔ IsSquare a := Iff.rfl +theorem mem_square {a : G} : a ∈ square G ↔ IsSquare a := Iff.rfl @[to_additive (attr := simp, norm_cast)] -theorem coe_squareIn : squareIn G = {s : G | IsSquare s} := rfl +theorem coe_square : square G = {s : G | IsSquare s} := rfl end Subgroup diff --git a/Mathlib/Algebra/Ring/SumsOfSquares.lean b/Mathlib/Algebra/Ring/SumsOfSquares.lean index 83f2856af82f4..64898b4df47db 100644 --- a/Mathlib/Algebra/Ring/SumsOfSquares.lean +++ b/Mathlib/Algebra/Ring/SumsOfSquares.lean @@ -21,15 +21,15 @@ 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 term `sumSq R` : in an additive monoid with multiplication `R`, we introduce +`sumSq 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`. - `IsSumSq.nonneg`: in a linearly ordered semiring `R` with an `[ExistsAddOfLE R]` instance, sums of squares are non-negative. -- `mem_sumSqIn_of_isSquare`: a square is a sum of squares. +- `mem_sumSq_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`. @@ -73,34 +73,36 @@ theorem IsSumSq.sum_mul_self {ι : Type*} [AddCommMonoid R] (s : Finset ι) (f : 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`, `AddSubmonoid.sumSq R` is the submonoid of sums of squares in `R`. -/ -def sumSqIn [AddMonoid R] : AddSubmonoid R where +def sumSq [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 +@[deprecated (since := "2024-08-09")] alias SumSqIn := sumSq +@[deprecated (since := "2025-01-03")] alias sumSqIn := sumSq /-- 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 mem_sumSqIn_of_isSquare [AddMonoid R] {x : R} (px : IsSquare x) : x ∈ sumSqIn R := by +theorem mem_sumSq_of_isSquare [AddMonoid R] {x : R} (px : IsSquare x) : x ∈ sumSq R := by rcases px with ⟨y, py⟩ rw [py, ← AddMonoid.add_zero (y * y)] exact IsSumSq.sq_add _ _ IsSumSq.zero -@[deprecated (since := "2024-08-09")] alias SquaresInSumSq := mem_sumSqIn_of_isSquare +@[deprecated (since := "2024-08-09")] alias SquaresInSumSq := mem_sumSq_of_isSquare +@[deprecated (since := "2025-01-03")] alias mem_sumSqIn_of_isSquare := mem_sumSq_of_isSquare /-- In an additive monoid with multiplication `R`, the submonoid generated by the squares is the set of sums of squares. -/ -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 ↦ ?_) +theorem AddSubmonoid.closure_isSquare [AddMonoid R] : closure {x : R | IsSquare x} = sumSq R := by + refine le_antisymm (closure_le.2 (fun x hx ↦ mem_sumSq_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